Presentation is loading. Please wait.

Presentation is loading. Please wait.

Semantics of Hoare Logic Aquinas Hobor and Martin Henz.

Similar presentations


Presentation on theme: "Semantics of Hoare Logic Aquinas Hobor and Martin Henz."— Presentation transcript:

1 Semantics of Hoare Logic Aquinas Hobor and Martin Henz

2 What does a Hoare tuple mean? { Á } P { Ã } Informal meaning (already given): “If the program P is run in a state that satisfies Á and P terminates, then the state resulting from P’s execution will satisfy Ã.” 2

3 We would like to formalize { Á } P { Ã } Informal meaning (already given): “If the program P is run in a state that satisfies Á and P terminates, then the state resulting from P’s execution will satisfy Ã.” 3

4 We would like to formalize { Á } P { Ã } Need to define: 1.Running a program P 2.P terminates 3.State satisfies Á 4.Resulting state satisfies Ã. 4

5 We would like to formalize { Á } P { Ã } For better clarity and “fun”, we will do it in Coq. 5

6 We would like to formalize { Á } P { Ã } For better clarity and “fun”, we will do it in Coq. (And by “we”, I mean I will do part of it in class and you will do the rest at home…) 6

7 We would like to formalize { Á } P { Ã } Need to define: 1.Running a program P 2.P terminates 3.State satisfies Á 4.Resulting state satisfies Ã. 7

8 Operational Semantics Numeric Expressions E: – n | x | (−E) | (E + E) | (E − E) | (E ∗ E) Boolean Expressions B: – true | false | (!B) | (B&B) | (B||B) | (E < E) Commands C: – x = E | C;C | if B {C} else {C} | while B {C} 8

9 We have to specify exactly how each evaluates Numeric Expressions E: – n | x | (−E) | (E + E) | (E − E) | (E ∗ E) First problem: what are our variables “x”? We will use our usual trick of letting variables be natural numbers: Definition var : Type := nat. 9

10 We have to specify exactly how each evaluates Numeric Expressions E: – n | x | (−E) | (E + E) | (E − E) | (E ∗ E) Next: how do we define our expressions? Definition num : Type := Z. Inductive nExpr : Type := | Num: forall n : num, nExpr | Var: forall v : var, nExpr | Neg: forall ne : nExpr, nExpr | Plus: forall ne1 ne2 : nExpr, nExpr | Minus: forall ne1 ne2 : nExpr, nExpr | Times: forall ne1 ne2 : nExpr, nExpr. 10

11 We have to specify exactly how each evaluates Numeric Expressions E: – n | x | (−E) | (E + E) | (E − E) | (E ∗ E) Now, what does evaluation of an E mean? We want to write E  n to mean “the expression E evaluates to the numeric n” But what about E = x? By itself, we don’t know what to do... 11

12 We have to specify exactly how each evaluates Numeric Expressions E: – n | x | (−E) | (E + E) | (E − E) | (E ∗ E) Define a context ° to be a function from variables to numbers. Definition ctx := var -> num. 12

13 We have to specify exactly how each evaluates Numeric Expressions E: – n | x | (−E) | (E + E) | (E − E) | (E ∗ E) Now define ° ` E  n to mean “in context °, the expression E evaluates to the numeric n.” 13

14 Numeric Evaluation in Coq Fixpoint neEval (g : ctx) (ne : nExpr) : num := match ne with | Num n => n | Var x => g x | Neg ne => 0 - (neEval g ne) | Plus ne1 ne2 => neEval g ne1) + (neEval g ne2) | Minus ne1 ne2 => (neEval g ne1) - (neEval g ne2) | Times ne1 ne2 => (neEval g ne1) * (neEval g ne2) end. 14

15 Boolean Expressions Boolean Expressions B: – true | false | (!B) | (B&B) | (B||B) | (E < E) Inductive bExpr : Type := | T : bExpr | F : bExpr | bNeg : forall be : bExpr, bExpr | And : forall be1 be2 : bExpr, bExpr | Or : forall be1 be2 : bExpr, bExpr | LT : forall ne1 ne2 : nExpr, bExpr. 15

16 Boolean Evaluation Boolean Expressions B: – true | false | (!B) | (B&B) | (B||B) | (E < E) Since B includes E, we will need contexts to evaluate Bs. What do we evaluate to? How about Prop. So define ° ` B  P to mean “in context °, the expression B evaluates to the proposition P.” 16

17 Boolean Evaluation Fixpoint beEval (g : ctx) (be : bExpr) : Prop := match be with | T => True | F => False | bNeg be => ~(beEval g be) | And be1 be2 => (beEval g be1) /\ (beEval g be2) | Or be1 be2 => (beEval g be1) \/ (beEval g be2) | LT ne1 ne2 => (neEval g ne1) < (neEval g ne2) end. 17

18 Commands Commands C: – x = E | C;C | if B {C} else {C} | while B {C} Inductive Coms : Type := | Assign : forall (x : var) (e: nExpr), Coms | Seq : forall c1 c2 : Coms, Coms | If : forall (b : bExpr) (c1 c2 : Coms), Coms | While : forall (b : bExpr) (c : Coms), Coms. 18

19 Command Evaluation Idea: executing command C for one step moves the machine from one state to the next What is a state ¾ ? Pair of context ° (data) and control k (code) Control k is either kStop (we are done) or kSeq C k – We can write C ² k for kSeq if that is easier – We can also write ¥ for kHalt 19

20 Controls and states in Coq Inductive ctl : Type := | kStop : ctl | kSeq : forall (c : Coms) (k : ctl), ctl. Definition state : Type := ctx * ctl. 20

21 Step relation, assign(a) We now define the step relation, written ¾ 1  ¾ 2 that is, “state ¾ 1 steps to state ¾ 2 ”, in parts: ° ` E  n ° ’ = [x ! n] ° ( °, (x = E) ² k)  ( ° ’, k) 21

22 Step relation, assign(b) ° ` E  n ° ’ = [x ! n] ° ( °, (x = E) ² k)  ( ° ’, k) | stepA: forall g x e n k g', n = neEval g e -> g' = upd_ctx g x n -> step (g, kSeq (Assign x e) k) (g', k) 22

23 Step relation, seq ( °, (C 1 ; C 2 ) ² k)  ( °, C 1 ² (C 2 ² k)) | stepS: forall g c1 c2 k, step (g, kSeq (Seq c1 c2) k) (g, kSeq c1 (kSeq c2 k)) 23

24 Step relation, if (a) ° ` B  True ( °, (if B then {C 1 } else {C 2 }) ² k)  ( °, C 1 ² k) | stepI1: forall g b c1 c2 k, beEval g b -> step (g, kSeq (If b c1 c2) k) (g, kSeq c1 k) 24

25 Step relation, if (b) ° ` B  False ( °, (if B then {C 1 } else {C 2 }) ² k)  ( °, C 2 ² k) | stepI2: forall g b c1 c2 k, ~beEval g b -> step (g, kSeq (If b c1 c2) k) (g, kSeq c2 k) 25

26 Step relation, while (a) ° ` B  True ( °, (while B {C}) ² k)  ( °, C ² (while B {C} ² k)) | stepW1: forall g b c k, beEval g b -> step (g, kSeq (While b c) k) (g, kSeq c (kSeq (While b c) k)) 26

27 Step relation, while (b) ° ` B  False ( °, (while B {C}) ² k)  ( °, k)) | stepW2: forall g b c k, ~beEval g b -> step (g, kSeq (While b c) k) (g, k) 27

28 Entire step relation ° ` E  n ° ’ = [x ! n] ° ( °, (x = E) ² k)  ( ° ’, k) ( °, (C 1 ; C 2 ) ² k)  ( °, C 1 ² (C 2 ² k)) ° ` B  True ( °, (if B then {C 1 } else {C 2 }) ² k)  ( °, C 1 ² k) ° ` B  False ( °, (if B then {C 1 } else {C 2 }) ² k)  ( °, C 2 ² k) ° ` B  True ( °, (while B {C}) ² k)  ( °, C ² (while B {C} ² k)) ° ` B  False ( °, (while B {C}) ² k)  ( °, k)) 28

29 Inductive step : state -> state -> Prop := | stepA: forall g x e k g', g' = upd_ctx g x (neEval g e) -> step (g, kSeq (Assign x e) k) (g', k) | stepS: forall g c1 c2 k, step (g, kSeq (Seq c1 c2) k) (g, kSeq c1 (kSeq c2 k)) | stepI1: forall g b c1 c2 k, beEval g b -> step (g, kSeq (If b c1 c2) k) (g, kSeq c1 k) | stepI2: forall g b c1 c2 k, ~beEval g b -> step (g, kSeq (If b c1 c2) k) (g, kSeq c2 k) | stepW1: forall g b c k, beEval g b -> step (g, kSeq (While b c) k) (g, kSeq c (kSeq (While b c) k)) | stepW2: forall g b c k, ~beEval g b -> step (g, kSeq (While b c) k) (g, k). 29

30 From step to step* Usually we want to run our program for more than one step. We write ¾  * ¾ ’ to mean that the state ¾ steps to the state ¾ ’ in some number of steps. 30

31 From step to step* ¾  * ¾ ¾  ¾ ’ ¾ ’  * ¾ ’’ ¾  * ¾ ’’ Inductive stepstar : state -> state -> Prop := | stepsZ: forall s, stepstar s s | stepsS: forall s s' s'', step s s' -> stepstar s' s'' -> stepstar s s''. 31

32 We would like to formalize { Á } P { Ã } Need to define: 1.Running a program P 2.P terminates 3.State satisfies Á 4.Resulting state satisfies Ã. 32

33 Halted; Terminates means eventually halted We say a state ( °, k) is halted when k = ¥ (* 20 points *) Lemma only_stuck_when_at_kStop: forall g k, (~ exists st', step (g, k) st') -> k = kStop. Proof. admit. Qed. ¾ terminates if 9 s’ such that ¾  * ¾ ’ and ¾ ’ is halted. 33

34 We would like to formalize { Á } P { Ã } Need to define: 1.Running a program P 2.P terminates 3.State satisfies Á 4.Resulting state satisfies Ã. 34

35 What is an assertion? We can do what we did for modal logic: Definition assertion : Type := ctx -> Prop. We can even write ° ² à as shorthand for à ( ° ) Thus, we can use the rules of our Coq metalogic to easily reason about our Hoare assertions. 35

36 Lifting Assertions to Metalogic 1 ° ² Á Æ Ã ´ ( ° ² Ã ) Æ ( ° ² Á ) Definition assertAnd (P Q : assertion) : assertion := fun g => P g /\ Q g. Notation "P && Q" := (assertAnd P Q). ° ² B ´ ° ` B  True Definition assertbEval (b : bExpr) : assertion := fun g => beEval g b. Notation "[ b ]" := (assertbEval b). 36

37 Lifting Assertions to Metalogic 2 ° ² [x ! e] Ã ´ [x ! n] ° ² Ã (where ° ` e  n) Definition assertReplace (x : var) (e : nExpr) (psi : assertion) : assertion := fun g => psi (upd_ctx g x (neEval g e)). Notation "[ x => e @ psi ]" := (assertReplace x e psi). ` AR Á ! Ã ´ 8 °, ( ° ² Á ) ) ( ° ² Ã ) Definition Implies (P Q: assertion) : Prop := forall g, P g -> Q g. Notation "P |-- Q" := (Implies P Q) (at level 30). 37

38 We would like to formalize { Á } P { Ã } Need to define: 1.Running a program P 2.P terminates 3.State satisfies Á 4.Resulting state satisfies Ã. 38

39 Putting it all together { Ã } C { Á } ´ 8 °. ° ² Ã ) 8 ° ’. ( °, C ² ¥ )  * ( ° ’, ¥ ) ) ° ’ ² Á Definition HTuple (pre : assertion) (c : Coms) (post : assertion) : Prop := forall g, pre g -> forall g', stepstar (g, kSeq c kStop) (g', kStop) -> post g'. 39

40 Now what? Prove the Hoare rules as lemmas from definitions! { Ã } c 1 { Â } { Â } c 2 { Á } { Ã } c 1 ; c 2 { Á } Lemma HT_Seq: forall a1 c1 a2 c2 a3, HTuple a1 c1 a2 -> HTuple a2 c2 a3 -> HTuple a1 (Seq c1 c2) a3. 40

41 Assignment Rule {[x ! E] Ã } x = E { Ã } Lemma HT_Asgn: forall x e psi, HTuple [x => e @ psi] (Assign x e) psi. 41

42 If Rule { Á Æ B} C 1 { Ã } { Á Æ : B} C 2 { Ã } { Á } if B {C 1 } else {C 2 } { Ã } Lemma HT_If: forall phi b c1 psi c2, HTuple (phi && [b]) c1 psi -> HTuple (phi && [bNeg b]) c2 psi -> HTuple phi (If b c1 c2) psi. 42

43 While Rule { Ã Æ B} C { Ã } { Ã } while B {C} { Ã Æ : B} Lemma HT_While: forall psi b c, HTuple (psi && [b]) c psi -> HTuple psi (While b c) (psi && [bNeg b]). 43

44 Implied Rule ` AR Á ’ ! Á { Á } C { Ã } ` AR Ã ! Ã ’ { Á ’} C { Ã ’} Lemma HT_Implied: forall phi phi' psi psi' c, phi' |-- phi -> HTuple phi c psi -> psi |-- psi' -> HTuple phi' c psi'. 44

45 Your task: Prove these lemmas HT_Seq : 10 points (with 20 extra credit if you solve the supporting lemma) HT_Asgn : 10 points HT_If : 10 points HT_Implied : 5 points HT_While : 40 points extra credit (good luck!) 45

46 Finally Definition x : var := 0. Definition y : var := 1. Definition z : var := 2. Open Local Scope Z_scope. Definition neq (ne1 ne2 : nExpr) : bExpr := Or (LT ne1 ne2) (LT ne2 ne1). Definition factorial_prog : Coms := Seq (Assign y (Num 1)) (* y := 1 *) (Seq (Assign z (Num 0)) (* z := 0 *) (While (neq (Var z) (Var x)) (* while z <> x { *) (Seq (Assign z (Plus (Var z) (Num 1))) (* z := z + 1 *) (Assign y (Times (Var y) (Var z)))(* y := y * z *) ) (* } *) ) ). 46

47 Statement of Theorem Definition Top : assertion := fun _ => True. Open Local Scope nat_scope. Fixpoint factorial (n : nat) := match n with | O => 1 | S n' => n * (factorial n') end. Open Local Scope Z_scope. Lemma factorial_good: HTuple Top factorial_prog (fun g => g y = Z_of_nat (factorial (Zabs_nat (g x)))). 47

48 Casts Definition Top : assertion := fun _ => True. Open Local Scope nat_scope. Fixpoint factorial (n : nat) := match n with | O => 1 | S n' => n * (factorial n') end. Open Local Scope Z_scope. Lemma factorial_good: HTuple Top factorial_prog (fun g => g y = Z_of_nat (factorial (Zabs_nat (g x)))). 48

49 Proof of Theorem Lemma factorial_good: HTuple Top factorial_prog (fun g => g y = Z_of_nat (factorial (Zabs_nat (g x)))). Proof. apply HT_Seq with (fun g => g y = 1). replace Top with ([y => (Num 1) @ (fun g : ctx => g y = 1)]). apply HT_Asgn. extensionality g. unfold assertReplace, Top, upd_ctx. simpl. apply prop_ext. firstorder. apply HT_Seq with (fun g :ctx => g z = 0 /\ g y = 1). replace (fun g : var -> Z => g y = 1) with ([z => (Num 0) @ (fun g :ctx => g z = 0 /\ g y = 1)]). apply HT_Asgn. extensionality g. unfold assertReplace, Top, upd_ctx. simpl. apply prop_ext. firstorder. apply HT_Implied with (fun g => g z >= 0 /\ g y = Z_of_nat (factorial (Zabs_nat (g z)))) ((fun g => g z >= 0 /\ g y = Z_of_nat (factorial (Zabs_nat (g z)))) && [bNeg (neq (Var z) (Var x))]). repeat intro. destruct H. rewrite H, H0. simpl. firstorder. apply HT_While. apply HT_Implied with (fun g => g z >=0 /\ (g y) * ((g z) + 1) = Z_of_nat (factorial (Zabs_nat ((g z) + 1)))) (fun g : ctx => g z - 1 >= 0 /\ g y = Z_of_nat (factorial (Zabs_nat (g z)))). repeat intro. destruct H. clear H0. rewrite H1. split; auto. remember (g z) as n. clear -H. destruct n; auto. simpl. rewrite <- Pplus_one_succ_r. rewrite nat_of_P_succ_morphism. simpl. remember (factorial (nat_of_P p)). clear. rewrite Zpos_succ_morphism. rewrite inj_plus. rewrite inj_mult. rewrite <- Zpos_eq_Z_of_nat_o_nat_of_P. ring. elimtype False. auto with zarith. apply HT_Seq with (fun g => g z - 1 >= 0 /\ g y * g z = Z_of_nat (factorial (Zabs_nat (g z)))). replace (fun g : var -> Z => g z >= 0 /\ g y * (g z + 1) = Z_of_nat (factorial (Zabs_nat (g z + 1)))) with [z => (Plus (Var z) (Num 1)) @ (fun g : var -> Z => g z - 1 >= 0 /\ g y * g z = Z_of_nat (factorial (Zabs_nat (g z))))]. apply HT_Asgn. extensionality g. apply prop_ext. firstorder. unfold upd_ctx in H. simpl in H. auto with zarith. simpl. unfold upd_ctx. simpl. auto with zarith. replace (fun g : var -> Z => g z - 1 >= 0 /\ g y * g z = Z_of_nat (factorial (Zabs_nat (g z)))) with [y => (Times (Var y) (Var z)) @ (fun g : var -> Z => g z - 1>= 0 /\ g y = Z_of_nat (factorial (Zabs_nat (g z))))]. apply HT_Asgn. extensionality g. apply prop_ext. firstorder. repeat intro; firstorder. repeat intro. destruct H. rewrite H1. simpl in H0. destruct (Ztrichotomy (g z) (g x)). contradiction H0; auto. destruct H2. rewrite <- H2. trivial. contradiction H0. right. apply Zgt_lt. trivial. Qed. 49

50 The good news… Your HW does not require you to do one of these yourself (we are not without mercy…) Still… why did I show it to you? 50

51 Seems like a lot of work… why bother? Lemma factorial_good: HTuple Top factorial_prog (fun g => g y = Z_of_nat (factorial (Zabs_nat (g x)))). Proof. apply HT_Seq with (fun g => g y = 1). replace Top with ([y => (Num 1) @ (fun g : ctx => g y = 1)]). apply HT_Asgn. extensionality g. unfold assertReplace, Top, upd_ctx. simpl. apply prop_ext. firstorder. apply HT_Seq with (fun g :ctx => g z = 0 /\ g y = 1). replace (fun g : var -> Z => g y = 1) with ([z => (Num 0) @ (fun g :ctx => g z = 0 /\ g y = 1)]). apply HT_Asgn. extensionality g. unfold assertReplace, Top, upd_ctx. simpl. apply prop_ext. firstorder. apply HT_Implied with (fun g => g z >= 0 /\ g y = Z_of_nat (factorial (Zabs_nat (g z)))) ((fun g => g z >= 0 /\ g y = Z_of_nat (factorial (Zabs_nat (g z)))) && [bNeg (neq (Var z) (Var x))]). repeat intro. destruct H. rewrite H, H0. simpl. firstorder. apply HT_While. apply HT_Implied with (fun g => g z >=0 /\ (g y) * ((g z) + 1) = Z_of_nat (factorial (Zabs_nat ((g z) + 1)))) (fun g : ctx => g z - 1 >= 0 /\ g y = Z_of_nat (factorial (Zabs_nat (g z)))). repeat intro. destruct H. clear H0. rewrite H1. split; auto. remember (g z) as n. clear -H. destruct n; auto. simpl. rewrite <- Pplus_one_succ_r. rewrite nat_of_P_succ_morphism. simpl. remember (factorial (nat_of_P p)). clear. rewrite Zpos_succ_morphism. rewrite inj_plus. rewrite inj_mult. rewrite <- Zpos_eq_Z_of_nat_o_nat_of_P. ring. elimtype False. auto with zarith. apply HT_Seq with (fun g => g z - 1 >= 0 /\ g y * g z = Z_of_nat (factorial (Zabs_nat (g z)))). replace (fun g : var -> Z => g z >= 0 /\ g y * (g z + 1) = Z_of_nat (factorial (Zabs_nat (g z + 1)))) with [z => (Plus (Var z) (Num 1)) @ (fun g : var -> Z => g z - 1 >= 0 /\ g y * g z = Z_of_nat (factorial (Zabs_nat (g z))))]. apply HT_Asgn. extensionality g. apply prop_ext. firstorder. unfold upd_ctx in H. simpl in H. auto with zarith. simpl. unfold upd_ctx. simpl. auto with zarith. replace (fun g : var -> Z => g z - 1 >= 0 /\ g y * g z = Z_of_nat (factorial (Zabs_nat (g z)))) with [y => (Times (Var y) (Var z)) @ (fun g : var -> Z => g z - 1>= 0 /\ g y = Z_of_nat (factorial (Zabs_nat (g z))))]. apply HT_Asgn. extensionality g. apply prop_ext. firstorder. repeat intro; firstorder. repeat intro. destruct H. rewrite H1. simpl in H0. destruct (Ztrichotomy (g z) (g x)). contradiction H0; auto. destruct H2. rewrite <- H2. trivial. contradiction H0. right. apply Zgt_lt. trivial. Qed. 51

52 Bug in Paper Proof Lemma factorial_good: HTuple Top factorial_prog (fun g => g y = Z_of_nat (factorial (Zabs_nat (g x)))). Proof. apply HT_Seq with (fun g => g y = 1). replace Top with ([y => (Num 1) @ (fun g : ctx => g y = 1)]). apply HT_Asgn. extensionality g. unfold assertReplace, Top, upd_ctx. simpl. apply prop_ext. firstorder. apply HT_Seq with (fun g :ctx => g z = 0 /\ g y = 1). replace (fun g : var -> Z => g y = 1) with ([z => (Num 0) @ (fun g :ctx => g z = 0 /\ g y = 1)]). apply HT_Asgn. extensionality g. unfold assertReplace, Top, upd_ctx. simpl. apply prop_ext. firstorder. apply HT_Implied with (fun g => g z >= 0 /\ g y = Z_of_nat (factorial (Zabs_nat (g z)))) ((fun g => g z >= 0 /\ g y = Z_of_nat (factorial (Zabs_nat (g z)))) && [bNeg (neq (Var z) (Var x))]). repeat intro. destruct H. rewrite H, H0. simpl. firstorder. apply HT_While. apply HT_Implied with (fun g => g z >=0 /\ (g y) * ((g z) + 1) = Z_of_nat (factorial (Zabs_nat ((g z) + 1)))) (fun g : ctx => g z - 1 >= 0 /\ g y = Z_of_nat (factorial (Zabs_nat (g z)))). repeat intro. destruct H. clear H0. rewrite H1. split; auto. remember (g z) as n. clear -H. destruct n; auto. simpl. rewrite <- Pplus_one_succ_r. rewrite nat_of_P_succ_morphism. simpl. remember (factorial (nat_of_P p)). clear. rewrite Zpos_succ_morphism. rewrite inj_plus. rewrite inj_mult. rewrite <- Zpos_eq_Z_of_nat_o_nat_of_P. ring. elimtype False. auto with zarith. apply HT_Seq with (fun g => g z - 1 >= 0 /\ g y * g z = Z_of_nat (factorial (Zabs_nat (g z)))). replace (fun g : var -> Z => g z >= 0 /\ g y * (g z + 1) = Z_of_nat (factorial (Zabs_nat (g z + 1)))) with [z => (Plus (Var z) (Num 1)) @ (fun g : var -> Z => g z - 1 >= 0 /\ g y * g z = Z_of_nat (factorial (Zabs_nat (g z))))]. apply HT_Asgn. extensionality g. apply prop_ext. firstorder. unfold upd_ctx in H. simpl in H. auto with zarith. simpl. unfold upd_ctx. simpl. auto with zarith. replace (fun g : var -> Z => g z - 1 >= 0 /\ g y * g z = Z_of_nat (factorial (Zabs_nat (g z)))) with [y => (Times (Var y) (Var z)) @ (fun g : var -> Z => g z - 1>= 0 /\ g y = Z_of_nat (factorial (Zabs_nat (g z))))]. apply HT_Asgn. extensionality g. apply prop_ext. firstorder. repeat intro; firstorder. repeat intro. destruct H. rewrite H1. simpl in H0. destruct (Ztrichotomy (g z) (g x)). contradiction H0; auto. destruct H2. rewrite <- H2. trivial. contradiction H0. right. apply Zgt_lt. trivial. Qed. 52

53 Forgot to track boundary condition (z >= 0 at all times in the loop) Lemma factorial_good: HTuple Top factorial_prog (fun g => g y = Z_of_nat (factorial (Zabs_nat (g x)))). Proof. apply HT_Seq with (fun g => g y = 1). replace Top with ([y => (Num 1) @ (fun g : ctx => g y = 1)]). apply HT_Asgn. extensionality g. unfold assertReplace, Top, upd_ctx. simpl. apply prop_ext. firstorder. apply HT_Seq with (fun g :ctx => g z = 0 /\ g y = 1). replace (fun g : var -> Z => g y = 1) with ([z => (Num 0) @ (fun g :ctx => g z = 0 /\ g y = 1)]). apply HT_Asgn. extensionality g. unfold assertReplace, Top, upd_ctx. simpl. apply prop_ext. firstorder. apply HT_Implied with (fun g => g z >= 0 /\ g y = Z_of_nat (factorial (Zabs_nat (g z)))) ((fun g => g z >= 0 /\ g y = Z_of_nat (factorial (Zabs_nat (g z)))) && [bNeg (neq (Var z) (Var x))]). repeat intro. destruct H. rewrite H, H0. simpl. firstorder. apply HT_While. apply HT_Implied with (fun g => g z >=0 /\ (g y) * ((g z) + 1) = Z_of_nat (factorial (Zabs_nat ((g z) + 1)))) (fun g : ctx => g z - 1 >= 0 /\ g y = Z_of_nat (factorial (Zabs_nat (g z)))). repeat intro. destruct H. clear H0. rewrite H1. split; auto. remember (g z) as n. clear -H. destruct n; auto. simpl. rewrite <- Pplus_one_succ_r. rewrite nat_of_P_succ_morphism. simpl. remember (factorial (nat_of_P p)). clear. rewrite Zpos_succ_morphism. rewrite inj_plus. rewrite inj_mult. rewrite <- Zpos_eq_Z_of_nat_o_nat_of_P. ring. elimtype False. auto with zarith. apply HT_Seq with (fun g => g z - 1 >= 0 /\ g y * g z = Z_of_nat (factorial (Zabs_nat (g z)))). replace (fun g : var -> Z => g z >= 0 /\ g y * (g z + 1) = Z_of_nat (factorial (Zabs_nat (g z + 1)))) with [z => (Plus (Var z) (Num 1)) @ (fun g : var -> Z => g z - 1 >= 0 /\ g y * g z = Z_of_nat (factorial (Zabs_nat (g z))))]. apply HT_Asgn. extensionality g. apply prop_ext. firstorder. unfold upd_ctx in H. simpl in H. auto with zarith. simpl. unfold upd_ctx. simpl. auto with zarith. replace (fun g : var -> Z => g z - 1 >= 0 /\ g y * g z = Z_of_nat (factorial (Zabs_nat (g z)))) with [y => (Times (Var y) (Var z)) @ (fun g : var -> Z => g z - 1>= 0 /\ g y = Z_of_nat (factorial (Zabs_nat (g z))))]. apply HT_Asgn. extensionality g. apply prop_ext. firstorder. repeat intro; firstorder. repeat intro. destruct H. rewrite H1. simpl in H0. destruct (Ztrichotomy (g z) (g x)). contradiction H0; auto. destruct H2. rewrite <- H2. trivial. contradiction H0. right. apply Zgt_lt. trivial. Qed. 53

54 Coercions (easily forgotten about…) Fixpoint factorial (n : nat) := match n with | O => 1 | S n' => n * (factorial n') end. fun g => g y = Z_of_nat (factorial (Zabs_nat (g x)))). We define factorial on nats because that way we have the best chance of not making a mistake in our specification. But there is a cost: we must coerce from Z to N and back to Z… 54

55 Where you need this fact in the proof Our “x!” has an implicit coercion in it: first we take the integer x, get the absolute value of it, and then calculate factorial on nats (and then coerce back to Z)… while (z <> x) { {y = z! Æ z <> x} Now use Implied {y * (z + 1) = (z + 1)!} 55

56 Where you need this fact in the proof Our “x!” has an implicit coercion in it: first we take the integer x, get the absolute value of it, and then calculate factorial on nats (and then coerce back to Z)… while (z <> x) { {y = z! Æ z <> x} Now use Implied {y * (z + 1) = (z + 1)!}  But wait! What if z < 0? Try y = 3, z = -4: 3 * (-4 + 1) = -9 (-4 + 1)! = (-3)! = 3! = 6 56

57 The Explosion of the Ariane 5 On June 4, 1996 an unmanned Ariane 5 rocket launched by the European Space Agency exploded just forty seconds after its lift- off from Kourou, French Guiana. The rocket was on its first voyage, after a decade of development costing $7 billion. The destroyed rocket and its cargo were valued at $500 million. A board of inquiry investigated the causes of the explosion and in two weeks issued a report. It turned out that the cause of the failure was a software error in the inertial reference system. Specifically a 64 bit floating point number relating to the horizontal velocity of the rocket with respect to the platform was converted to a 16 bit signed integer. The number was larger than 32,767, the largest integer storable in a 16 bit signed integer, and thus the conversion failed. 57

58 58


Download ppt "Semantics of Hoare Logic Aquinas Hobor and Martin Henz."

Similar presentations


Ads by Google