diff --git a/metadata/entries/HyperHoareLogic.toml b/metadata/entries/HyperHoareLogic.toml --- a/metadata/entries/HyperHoareLogic.toml +++ b/metadata/entries/HyperHoareLogic.toml @@ -1,31 +1,33 @@ title = "Formalization of Hyper Hoare Logic: A Logic to (Dis-)Prove Program Hyperproperties" date = 2023-04-03 topics = [ "Computer science/Programming languages/Logics", ] abstract = """ -Hoare logics are proof systems that allow one to formally establish properties of computer programs. Traditional Hoare logics prove properties of individual program executions (so-called trace properties, such as functional correctness). On the one hand, Hoare logic has been generalized to prove properties of multiple executions of a program (so-called hyperproperties, such as determinism or non-interference). These program logics prove the absence of (bad combinations of) executions. On the other hand, program logics similar to Hoare logic have been proposed to disprove program properties (e.g., Incorrectness Logic [8]), by proving the existence of (bad combinations of) executions. All of these logics have in common that they specify program properties using assertions over a fixed number of states, for instance, a single pre- and post-state for functional properties or pairs of pre- and post-states for non-interference. +Hoare logics are proof systems that allow one to formally establish properties of computer programs. Traditional Hoare logics prove properties of individual program executions (such as functional correctness). On the one hand, Hoare logic has been generalized to prove properties of multiple executions of a program (so-called hyperproperties, such as determinism or non-interference). These program logics prove the absence of (bad combinations of) executions. On the other hand, program logics similar to Hoare logic have been proposed to disprove program properties (e.g., Incorrectness Logic), by proving the existence of (bad combinations of) executions. All of these logics have in common that they specify program properties using assertions over a fixed number of states, for instance, a single pre- and post-state for functional properties or pairs of pre- and post-states for non-interference. -In this entry, we formalize Hyper Hoare Logic, a generalization of Hoare logic that lifts assertions to properties of arbitrary sets of states. The resulting logic is simple yet expressive: its judgments can express arbitrary trace- and hyperproperties over the terminating executions of a program. By allowing assertions to reason about sets of states, Hyper Hoare Logic can reason about both the absence and the existence of (combinations of) executions, and, thereby, supports both proving and disproving program (hyper-)properties within the same logic. In fact, we prove that Hyper Hoare Logic subsumes the properties handled by numerous existing correctness and incorrectness logics, and can express hyperproperties that no existing Hoare logic can. We also prove that Hyper Hoare Logic is sound and complete.""" +In this entry, we formalize Hyper Hoare Logic, a generalization of Hoare logic that lifts assertions to properties of arbitrary sets of states. The resulting logic is simple yet expressive: its judgments can express arbitrary program hyperproperties, a particular class of hyperproperties over the set of terminating executions of a program (including properties of individual program executions). By allowing assertions to reason about sets of states, Hyper Hoare Logic can reason about both the absence and the existence of (combinations of) executions, and, thereby, supports both proving and disproving program (hyper-)properties within the same logic, including hyperproperties that no existing Hoare logic can express. We prove that Hyper Hoare Logic is sound and complete, and demonstrate that it captures important proof principles naturally.""" license = "bsd" note = "" [authors] [authors.dardinier] email = "dardinier_email" [contributors] [notify] dardinier = "dardinier_email" [history] +<2024>-<04>-<12> = "additional rules for compositionality, syntactic hyper-assertions, and loops" [extra] [related] dois = [ "10.48550/arXiv.2301.10037", + "10.1145/3656437" ] pubs = [] diff --git a/thys/HyperHoareLogic/Compositionality.thy b/thys/HyperHoareLogic/Compositionality.thy new file mode 100644 --- /dev/null +++ b/thys/HyperHoareLogic/Compositionality.thy @@ -0,0 +1,908 @@ +section \Compositionality Rules\ + +theory Compositionality + imports Logic Expressivity Loops +begin + +text \In this file, we prove the soundness of all compositionality rules presented in Appendix D (figure 11).\ + +definition in_set where + "in_set \ S \ \ \ S" + +subsection \Linking rule\ + +proposition rule_linking: + assumes "\\1 (\2 :: ('a, 'b, 'c, 'd) state). fst \1 = fst \2 \ ( \ { (in_set \1 :: (('a, 'b, 'c, 'd) state) hyperassertion) } C { in_set \2 } ) + \ ( \ { (P \1 :: (('a, 'b, 'c, 'd) state) hyperassertion) } C { Q \2 } )" + shows "\ { ((\S. \\1 \ S. P \1 S) :: (('a, 'b, 'c, 'd) state) hyperassertion) } C { (\S. \\2 \ S. Q \2 S) }" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "\\1\S. P \1 S" + show "\\2\sem C S. Q \2 (sem C S)" + proof (clarify) + fix l \' assume "(l, \') \ sem C S" + then obtain \ where "(l, \) \ S" "single_sem C \ \'" + by (metis fst_conv in_sem snd_conv) + then show "Q (l, \') (sem C S)" + by (metis (mono_tags, opaque_lifting) asm0 assms fst_eqD hyper_hoare_triple_def in_set_def single_step_then_in_sem) + qed +qed + + +lemma rule_linking_alt: + assumes "\l \ \'. single_sem C \ \' \ ( \ { P (l, \) } C { Q (l, \') } )" + shows "\ { (\S. \\ \ S. P \ S) } C { (\S. \\' \ S. Q \' S) }" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "\\\S. P \ S" + show "\\'\sem C S. Q \' (sem C S)" + proof (clarify) + fix l \' assume "(l, \') \ sem C S" + then obtain \ where "(l, \) \ S" "single_sem C \ \'" + by (metis fst_conv in_sem snd_conv) + then have "\ { P (l, \) } C { Q (l, \') }" + by (simp add: assms) + then show "Q (l, \') (sem C S)" + by (simp add: \(l, \) \ S\ asm0 hyper_hoare_tripleE) + qed +qed + +subsection \Frame rules\ + +lemma rule_lframe: + + fixes b :: "('a \ ('lvar \ 'lval)) \ bool" +\\b takes a mapping from keys to logical states (representing the tuple), and returns a boolean\ + + shows "\ { (\S. \\. (\k. \ k \ S) \ b (fst \ \)) } C { \S. \\. (\k. \ k \ S) \ b (fst \ \) }" + +proof (rule hyper_hoare_tripleI) + fix S :: "('lvar, 'lval, 'b, 'c) state set" + assume asm0: "\\. (\k. \ k \ S) \ b (fst \ \)" + show "\\. (\k. \ k \ sem C S) \ b (fst \ \)" + proof (clarify) + fix \' :: "'a \ ('lvar, 'lval, 'b, 'c) state" + assume asm1: "\k. \' k \ sem C S" + let ?\ = "\k. SOME \k. fst \k = fst (\' k) \ \k \ S \ single_sem C (snd \k) (snd (\' k))" + have r: "\k. fst (?\ k) = fst (\' k) \ (?\ k) \ S \ single_sem C (snd (?\ k)) (snd (\' k))" + proof - + fix k show "fst (?\ k) = fst (\' k) \ (?\ k) \ S \ single_sem C (snd (?\ k)) (snd (\' k))" + proof (rule someI_ex) + show "\x. fst x = fst (\' k) \ x \ S \ \C, snd x\ \ snd (\' k)" + by (metis asm1 fst_conv in_sem snd_conv) + qed + qed + then have "b (fst \ ?\)" + using asm0 by presburger + moreover have "fst \ ?\ = fst \ \'" + using ext r by fastforce + then show "b (fst \ \')" + using calculation by presburger + qed +qed + +lemma rule_lframe_single: + "\ { (\S. \\ \ S. P (fst \)) } C { \S. \\ \ S. P (fst \) }" +proof - + let ?P = "\\. P (\ ())" + have "\ { (\S. \\. (\k. \ k \ S) \ ?P (fst \ \)) } C { \S. \\. (\k. \ k \ S) \ ?P (fst \ \) }" + using rule_lframe by fast + moreover have "(\S. \\ \ S. P (fst \)) = (\S. \\. (\k. \ k \ S) \ ?P (fst \ \))" + using ext by fastforce + ultimately show ?thesis + using forw_subst ssubst order_trans_rules(13) prop_subst basic_trans_rules(13) + by fast +qed + +definition differ_only_by_pset where + "differ_only_by_pset vars a b \ (\i. fst (a i) = fst (b i) \ differ_only_by_set vars (snd (a i)) (snd (b i)))" + +lemma differ_only_by_psetI: + assumes "\i. fst (a i) = fst (b i) \ differ_only_by_set vars (snd (a i)) (snd (b i))" + shows "differ_only_by_pset vars a b" + by (simp add: assms differ_only_by_pset_def) + +definition not_in_free_pvars_prop where + "not_in_free_pvars_prop vars b \ (\\1 \2. differ_only_by_pset vars \1 \2 \ (b \1 \ b \2))" + +proposition rule_frame: + + fixes b :: "('a \ ('lvar, 'lval, 'pvar, 'pval) state) \ bool" +\\b takes a mapping from keys to extended states (representing the tuple), and returns a boolean\ + + assumes "not_in_free_pvars_prop (written_vars C) b" + + shows "\ { (\S. \\. (\k. \ k \ S) \ b \) } C { \S. \\. (\k. \ k \ S) \ b \ }" + +proof (rule hyper_hoare_tripleI) + fix S :: "('lvar, 'lval, 'pvar, 'pval) state set" + assume asm0: "\\. (\k. \ k \ S) \ b \" + show "\\. (\k. \ k \ sem C S) \ b \" + proof (clarify) + fix \' :: "'a \ ('lvar, 'lval, 'pvar, 'pval) state" + assume asm1: "\k. \' k \ sem C S" + let ?\ = "\k. SOME \k. fst \k = fst (\' k) \ \k \ S \ single_sem C (snd \k) (snd (\' k))" + have r: "\k. fst (?\ k) = fst (\' k) \ (?\ k) \ S \ single_sem C (snd (?\ k)) (snd (\' k))" + proof - + fix k show "fst (?\ k) = fst (\' k) \ (?\ k) \ S \ single_sem C (snd (?\ k)) (snd (\' k))" + proof (rule someI_ex) + show "\x. fst x = fst (\' k) \ x \ S \ \C, snd x\ \ snd (\' k)" + by (metis asm1 fst_conv in_sem snd_conv) + qed + qed + then have "b ?\" + using asm0 by presburger + moreover have "differ_only_by_pset (written_vars C) ?\ \'" + proof (rule differ_only_by_psetI) + fix i show "fst (SOME \k. fst \k = fst (\' i) \ \k \ S \ \C, snd \k\ \ snd (\' i)) = fst (\' i) \ + differ_only_by_set (written_vars C) (snd (SOME \k. fst \k = fst (\' i) \ \k \ S \ \C, snd \k\ \ snd (\' i))) (snd (\' i))" + by (metis (mono_tags, lifting) differ_only_by_set_def r written_vars_not_modified) + qed + ultimately show "b \'" + using assms not_in_free_pvars_prop_def by blast + qed +qed + + +subsection \Logical Updates\ + +definition equal_outside_set where + "equal_outside_set vars l1 l2 \ (\x. x \ vars \ l1 x = l2 x)" + +lemma equal_outside_setI: + assumes "\x. x \ vars \ l1 x = l2 x" + shows "equal_outside_set vars l1 l2" + using assms equal_outside_set_def by auto + +lemma equal_outside_setE: + assumes "equal_outside_set vars l1 l2" + and "x \ vars" + shows "l1 x = l2 x" + using assms equal_outside_set_def by meson + +lemma equal_outside_sym: + "equal_outside_set vars l l' \ equal_outside_set vars l' l" + by (metis equal_outside_set_def) + +definition subset_mod_updates where + "subset_mod_updates vars S S' \ (\\ \ S. \\' \ S'. snd \ = snd \' \ equal_outside_set vars (fst \) (fst \'))" + +lemma subset_mod_updatesI: + assumes "\\. \ \ S \ (\\' \ S'. snd \ = snd \' \ equal_outside_set vars (fst \) (fst \'))" + shows "subset_mod_updates vars S S'" + by (simp add: assms subset_mod_updates_def) + +lemma subset_mod_updatesE: + assumes "subset_mod_updates vars S S'" + and "\ \ S" + shows "\\' \ S'. snd \ = snd \' \ equal_outside_set vars (fst \) (fst \')" + using assms(1) assms(2) subset_mod_updates_def by blast + + +definition same_mod_updates where + "same_mod_updates vars S S' \ subset_mod_updates vars S S' \ subset_mod_updates vars S' S" + +lemma same_mod_updatesI: + assumes "\\. \ \ S \ (\\' \ S'. snd \ = snd \' \ equal_outside_set vars (fst \) (fst \'))" + and "\\'. \' \ S' \ (\\ \ S. snd \ = snd \' \ equal_outside_set vars (fst \') (fst \))" + shows "same_mod_updates vars S S'" + by (metis assms(1) assms(2) same_mod_updates_def subset_mod_updates_def) + +lemma same_mod_updates_sym: + "same_mod_updates vars S S' \ same_mod_updates vars S' S" + using same_mod_updates_def by blast + +lemma same_mod_updates_refl: + "same_mod_updates vars S S" + by (metis equal_outside_set_def same_mod_updates_def subset_mod_updates_def) + +lemma equal_outside_set_trans: + assumes "equal_outside_set vars a b" + and "equal_outside_set vars b c" + shows "equal_outside_set vars a c" + using equal_outside_set_def[of vars a b] equal_outside_set_def[of vars b c] equal_outside_set_def[of vars a c] assms + by presburger + +lemma subset_mod_updates_trans: + assumes "subset_mod_updates vars S1 S2" + and "subset_mod_updates vars S2 S3" + shows "subset_mod_updates vars S1 S3" +proof (rule subset_mod_updatesI) + fix \1 assume "\1 \ S1" + then obtain \2 where "\2 \ S2" "snd \1 = snd \2 \ equal_outside_set vars (fst \1) (fst \2)" + using assms(1) same_mod_updates_def subset_mod_updatesE by blast + then show "\\'\S3. snd \1 = snd \' \ equal_outside_set vars (fst \1) (fst \')" + by (metis (no_types, lifting) assms(2) equal_outside_set_trans subset_mod_updatesE) +qed + +lemma same_mod_updates_trans: + assumes "same_mod_updates vars S1 S2" + and "same_mod_updates vars S2 S3" + shows "same_mod_updates vars S1 S3" + by (meson assms(1) assms(2) same_mod_updates_def subset_mod_updates_trans) + + +lemma sem_update_commute_aux: + assumes "subset_mod_updates vars S1 S2" + shows "subset_mod_updates vars (sem C S1) (sem C S2)" +proof (rule subset_mod_updatesI) + fix \1 assume asm0: "\1 \ sem C S1" + then obtain l1 \ where "(l1, \) \ S1" "single_sem C \ (snd \1)" "fst \1 = l1" + by (metis in_sem) + then obtain l2 where "(l2, \) \ S2" "equal_outside_set vars l1 l2" + using assms subset_mod_updatesE by fastforce + then show "\\'\sem C S2. snd \1 = snd \' \ equal_outside_set vars (fst \1) (fst \')" + by (metis \\C, \\ \ snd \1\ \fst \1 = l1\ fst_conv in_sem snd_conv) +qed + + +lemma sem_update_commute: + assumes "same_mod_updates (vars :: 'a set) S1 S2" + shows "same_mod_updates vars (sem C S1) (sem C S2)" + using assms same_mod_updates_def sem_update_commute_aux by blast + + + +type_synonym ('a, 'b, 'c, 'd) chyperassertion = "(('a \ 'b) \ ('c \ 'd)) set \ bool" + +definition invariant_on_updates :: "'a set \ ('a, 'b, 'c, 'd) chyperassertion \ bool" where + "invariant_on_updates vars P \ (\S S'. same_mod_updates vars S S' \ (P S \ P S'))" + +lemma invariant_on_updatesI: + assumes "\S S'. same_mod_updates vars S S' \ P S \ P S'" + shows "invariant_on_updates vars P" + using assms invariant_on_updates_def same_mod_updates_sym by blast + +definition entails_with_updates :: "'a set \ ('a, 'b, 'c, 'd) chyperassertion \ ('a, 'b, 'c, 'd) chyperassertion \ bool" + where + "entails_with_updates vars P Q \ (\S. P S \ (\S'. same_mod_updates vars S S' \ Q S'))" + +lemma entails_with_updatesI: + assumes "\S. P S \ (\S'. same_mod_updates vars S S' \ Q S')" + shows "entails_with_updates vars P Q" + by (simp add: assms entails_with_updates_def) + + +lemma entails_with_updatesE: + assumes "entails_with_updates vars P Q" + and "P S" + shows "\S'. same_mod_updates vars S S' \ Q S'" + by (meson assms(1) assms(2) entails_with_updates_def) + + +proposition rule_LUpdate: + assumes "\ {P'} C {Q}" + and "entails_with_updates vars P P'" + and "invariant_on_updates vars Q" + shows "\ {P} C {Q}" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "P S" + then obtain S' where "same_mod_updates vars S S' \ P' S'" + using assms(2) entails_with_updatesE by blast + then have "same_mod_updates vars (sem C S) (sem C S')" + using sem_update_commute by blast + moreover have "Q (sem C S')" + using \same_mod_updates vars S S' \ P' S'\ assms(1) hyper_hoare_tripleE by auto + ultimately show "Q (sem C S)" + using assms(3) invariant_on_updates_def by blast +qed + + + + +subsection \Filters\ + +lemma filter_prop_commute_aux: + assumes "\\ \'. fst \ = fst \' \ (f \ \ f \')" + shows "Set.filter f (sem C S) = sem C (Set.filter f S)" (is "?A = ?B") +proof + show "?A \ ?B" + proof + fix \' assume "\' \ ?A" + then obtain \ where "\ \ S" "single_sem C (snd \) (snd \')" "fst \ = fst \'" "f \'" + by (metis fst_conv in_sem member_filter snd_conv) + then show "\' \ ?B" + by (metis assms in_sem member_filter prod.collapse) + qed + show "?B \ ?A" + proof + fix x assume "x \ sem C (Set.filter f S)" + then show "x \ Set.filter f (sem C S)" + by (metis assms fst_conv in_sem member_filter) + qed +qed + +definition commute_with_sem where + "commute_with_sem f \ (\S C. f (sem C S) = sem C (f S))" + +lemma commute_with_semI: + assumes "\(S :: (('a \ 'b) \ ('c \ 'd)) set) C. f (sem C S) = sem C (f S)" + shows "commute_with_sem f" + by (simp add: assms commute_with_sem_def) + +lemma filter_prop_commute: + assumes "\\ \'. fst \ = fst \' \ (f \ \ f \')" + shows "commute_with_sem (Set.filter f)" + using assms commute_with_sem_def filter_prop_commute_aux by blast + +lemma rule_apply: + assumes "\ {P} C {Q}" + and "commute_with_sem f" + shows "\ {P \ f} C {Q \ f}" +proof (rule hyper_hoare_tripleI) + fix S assume "(P \ f) S" + then show "(Q \ f) (sem C S)" + by (metis assms(1) assms(2) commute_with_sem_def comp_apply hyper_hoare_tripleE) +qed + +definition apply_filter where + "apply_filter b P S \ P (Set.filter b S)" + +proposition rule_LFilter: + assumes "\ {P} C {Q}" + shows "\ { P \ (Set.filter (b \ fst)) } C { Q \ (Set.filter (b \ fst)) }" + by (simp add: assms filter_prop_commute rule_apply) + + +definition differ_only_by_pset_single where + "differ_only_by_pset_single vars a b \ (fst a = fst b \ differ_only_by_set vars (snd a) (snd b))" + +definition not_in_free_pvars_pexp where + "not_in_free_pvars_pexp vars b \ (\\1 \2. differ_only_by_set vars \1 \2 \ (b \1 \ b \2))" + +lemma single_sem_differ_by_written_vars: + assumes "single_sem C \ \'" + shows "differ_only_by_set (written_vars C) \ \'" + by (meson assms differ_only_by_set_def written_vars_not_modified) + +lemma single_sem_not_free_vars: + assumes "not_in_free_pvars_pexp (written_vars C) b" + and "single_sem C \ \'" + shows "b \ \ b \'" + using assms(1) assms(2) not_in_free_pvars_pexp_def single_sem_differ_by_written_vars by blast + + +proposition rule_PFilter: + assumes "\ {P} C {Q}" + and "not_in_free_pvars_pexp (written_vars C) b" + shows "\ { P \ (Set.filter (b \ snd)) } C { Q \ (Set.filter (b \ snd)) }" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "(P \ Set.filter (b \ snd)) S" + let ?S = "Set.filter (b \ snd) S" + have "Q (sem C ?S)" + using asm0 assms(1) hyper_hoare_tripleE by auto + moreover have "sem C ?S = Set.filter (b \ snd) (sem C S)" (is "?A = ?B") + proof + show "?B \ ?A" + proof + fix \' assume "\' \ Set.filter (b \ snd) (sem C S)" + then obtain \ where "\ \ S" "fst \ = fst \'" "single_sem C (snd \) (snd \')" "b (snd \')" + by (metis comp_apply fst_conv in_sem member_filter snd_conv) + then have "b (snd \)" + using single_sem_not_free_vars[of C b "snd \" "snd \'"] assms(2) + by simp + then show "\' \ ?A" + by (metis \\C, snd \\ \ snd \'\ \\ \ S\ \fst \ = fst \'\ comp_apply in_sem member_filter prod.collapse) + qed + show "?A \ ?B" + proof + fix \' assume "\' \ sem C (Set.filter (b \ snd) S)" + then obtain \ where "\ \ S" "fst \ = fst \'" "single_sem C (snd \) (snd \')" "b (snd \)" + by (metis (mono_tags, lifting) comp_apply fst_conv in_sem member_filter snd_conv) + then have "b (snd \')" + using assms(2) single_sem_not_free_vars by blast + then show "\' \ ?B" + by (metis \\C, snd \\ \ snd \'\ \\ \ S\ \fst \ = fst \'\ comp_apply member_filter prod.collapse single_step_then_in_sem) + qed + qed + ultimately show "(Q \ Set.filter (b \ snd)) (sem C S)" + by auto +qed + + + + +subsection \Other Compositionality Rules\ + +proposition rule_False: + "hyper_hoare_triple (\_. False) C Q" + by (simp add: hyper_hoare_triple_def) + +proposition rule_True: + "hyper_hoare_triple P C (\_. True)" + by (simp add: hyper_hoare_triple_def) + + +(* Other direction does not hold! *) +lemma sem_inter: + "sem C (S1 \ S2) \ sem C S1 \ sem C S2" + by (simp add: sem_monotonic) + + +proposition rule_Union: + assumes "\ {P} C {Q}" + and "hyper_hoare_triple P' C Q'" + shows "hyper_hoare_triple (join P P') C (join Q Q')" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "join P P' S" + then obtain S1 S2 where "S = S1 \ S2" "P S1" "P' S2" + by (metis join_def) + then have "sem C S = sem C S1 \ sem C S2" + using sem_union by auto + then show "join Q Q' (sem C S)" + by (metis \P S1\ \P' S2\ assms(1) assms(2) hyper_hoare_tripleE join_def) +qed + +proposition rule_IndexedUnion: + assumes "\x. \ {P x} C {Q x}" + shows "hyper_hoare_triple (general_join P) C (general_join Q)" +proof (rule hyper_hoare_tripleI) + fix S assume "general_join P S" + then obtain F where asm0: "S = (\x. F x)" "\x. P x (F x)" + by (meson general_join_def) + have "sem C S = (\x. sem C (F x))" + by (simp add: asm0(1) sem_split_general) + moreover have "\x. Q x (sem C (F x))" + using asm0(2) assms hyper_hoare_tripleE by blast + ultimately show "general_join Q (sem C S)" + by (metis general_join_def) +qed + +proposition rule_And: + assumes "\ {P} C {Q}" + and "hyper_hoare_triple P' C Q'" + shows "hyper_hoare_triple (conj P P') C (conj Q Q')" +proof (rule hyper_hoare_tripleI) + fix S assume "Logic.conj P P' S" + then show "Logic.conj Q Q' (sem C S)" + by (metis assms(1) assms(2) conj_def hyper_hoare_tripleE) +qed + +lemma rule_Forall: + assumes "\x. \ {P x} C {Q x}" + shows "hyper_hoare_triple (forall P) C (forall Q)" + by (metis assms forall_def hyper_hoare_triple_def) + +lemma rule_Or: + assumes "\ {P} C {Q}" + and "\ {P'} C {Q'}" + shows "hyper_hoare_triple (disj P P') C (disj Q Q')" + by (metis assms(1) assms(2) disj_def hyper_hoare_triple_def) + +corollary variant_if_rule: + assumes "hyper_hoare_triple P C1 Q" + and "hyper_hoare_triple P C2 Q" + and "closed_by_union Q" + shows "hyper_hoare_triple P (If C1 C2) Q" + by (metis assms(1) assms(2) assms(3) if_rule join_closed_by_union) + + +text \Simplifying the rule\ + +definition stable_by_infinite_union :: "'a hyperassertion \ bool" where + "stable_by_infinite_union I \ (\F. (\S \ F. I S) \ I (\S \ F. S))" + +lemma stable_by_infinite_unionE: + assumes "stable_by_infinite_union I" + and "\S. S \ F \ I S" + shows "I (\S \ F. S)" + using assms(1) assms(2) stable_by_infinite_union_def by blast + +lemma stable_by_union_and_constant_then_I: + assumes "\n. I n = I'" + and "stable_by_infinite_union I'" + shows "natural_partition I = I'" +proof (rule ext) + fix S show "natural_partition I S = I' S" + proof + show "I' S \ natural_partition I S" + proof - + assume "I' S" + show "natural_partition I S" + proof (rule natural_partitionI) + show "S = \ (range (\n. S))" + by simp + fix n show "I n S" + by (simp add: \I' S\ assms(1)) + qed + qed + assume asm0: "natural_partition I S" + then obtain F where "S = (\n. F n)" "\n. I n (F n)" + using natural_partitionE by blast + let ?F = "{F n |n. True}" + have "I' (\S\?F. S)" + using assms(2) + proof (rule stable_by_infinite_unionE[of I']) + fix S assume "S \ {F n |n. True}" + then show "I' S" + using \\n. I n (F n)\ assms(1) by force + qed + moreover have "(\S\?F. S) = S" + using \S = (\n. F n)\ by auto + ultimately show "I' S" by blast + qed +qed + +corollary simpler_rule_while: + assumes "hyper_hoare_triple I C I" + and "stable_by_infinite_union I" + shows "hyper_hoare_triple I (While C) I" +proof - + let ?I = "\n. I" + have "hyper_hoare_triple (?I 0) (While C) (natural_partition ?I)" + using while_rule[of ?I C] + by (simp add: assms(1) assms(2) stable_by_union_and_constant_then_I) + then show ?thesis + by (simp add: assms(2) stable_by_union_and_constant_then_I) +qed + + +lemma rule_and3: + assumes "\ {P1} C {Q1}" + and "\ {P2} C {Q2}" + and "\ {P3} C {Q3}" + shows "\ { conj P1 (conj P2 P3) } C { conj Q1 (conj Q2 Q3) }" + by (simp add: assms(1) assms(2) assms(3) rule_And) + +definition not_empty where + "not_empty S \ S \ {}" + +definition finite_not_empty where + "finite_not_empty S \ S \ {} \ finite S" + +definition update_logical where + "update_logical \ i v = ((fst \)(i := v), snd \)" + + +lemma single_sem_prop: + assumes "single_sem C (snd \) (snd \')" + and "fst \ = fst \'" + shows "\ { (\S. \ \ S) } C { (\S. \' \ S) }" + by (metis assms(1) assms(2) hyper_hoare_tripleI in_sem prod.collapse) + +lemma weaker_linking_rule: + assumes "\l \ \'. \ { (\S. (l, \) \ S) } C { (\S. (l, \') \ S) } \ ( \ { P (l, \) } C { Q (l, \') } )" + shows "\ { (\S. \\ \ S. P \ S) } C { (\S. \\' \ S. Q \' S) }" + by (simp add: assms rule_linking_alt single_sem_prop) + + +definition general_union :: "'a hyperassertion \ 'a hyperassertion" where + "general_union P S \ (\F. S = Union F \ (\S' \ F. P S'))" + +lemma general_unionI: + assumes "S = Union F" + and "\S'. S' \ F \ P S'" + shows "general_union P S" + using assms(1) assms(2) general_union_def by blast + +lemma general_unionE: + assumes "general_union P S" + obtains F where "S = Union F" "\S'. S' \ F \ P S'" + by (metis assms general_union_def) + + +(* Derived *) +proposition rule_BigUnion: + fixes P :: "((('a \ 'b) \ ('c \ 'd)) set \ bool)" + assumes "\ {P} C {Q}" + shows "\ {general_union P} C {general_union Q}" +proof (rule consequence_rule) + define PP :: "(('a \ 'b) \ ('c \ 'd)) set \ (('a \ 'b) \ ('c \ 'd)) set \ bool" where + "PP = (\_. P)" + let ?P = "disj (general_join PP) (\S. S = {})" + show "entails (general_union P) ?P" + proof (rule entailsI) + fix S assume "general_union P S" + then obtain F where "S = Union F" "\S'. S' \ F \ P S'" + by (metis general_union_def) + have "general_join PP S \ S = {}" + proof (cases "S = {}") + case True + then show ?thesis + by simp + next + case False + then obtain S' where "S' \ F" + using \S = \ F\ by blast + let ?F = "\SS. if SS \ F then SS else S'" + have "general_join PP S" + proof (rule general_joinI) + fix x show "PP x (?F x)" + by (simp add: PP_def \S' \ F\ \\S'. S' \ F \ P S'\) + next + show "S = (\x. if x \ F then x else S')" + using \S = \ F\ \S' \ F\ by force + qed + then show ?thesis by simp + qed + then show "disj (general_join PP) (\S. S = {}) S" + by (simp add: disj_def) + qed + + define QQ :: "(('a \ 'b) \ ('c \ 'd)) set \ (('a \ 'b) \ ('c \ 'd)) set \ bool" where + "QQ = (\_. Q)" + let ?Q = "disj (general_join QQ) (\S. S = {})" + + show "\ {Logic.disj (general_join PP) (\S. S = {})} C {?Q}" + proof (rule rule_Or) + show "\ {(\S. S = {})} C {\S. S = {}}" + by (metis (mono_tags, lifting) empty_iff equals0I hyper_hoare_triple_def in_sem) + (* TODO: Prove this rule *) + show "\ {general_join PP} C {general_join QQ}" + proof (rule rule_IndexedUnion) + fix x show "\ {PP x} C {QQ x}" + by (simp add: PP_def QQ_def assms) + qed + qed + show "entails (Logic.disj (general_join QQ) (\S. S = {})) (general_union Q)" + proof (rule entailsI) + fix S assume "Logic.disj (general_join QQ) (\S. S = {}) S" + then show "general_union Q S" + proof (cases "S = {}") + case True + then show ?thesis + using general_union_def by auto + next + case False + then obtain F where "\x. QQ x (F x)" "S = \ (range F)" + by (metis (mono_tags, lifting) \Logic.disj (general_join QQ) (\S. S = {}) S\ disj_def general_join_def) + then show ?thesis + by (metis QQ_def general_unionI rangeE) + qed + qed +qed + + + +proposition rule_Empty: + "\ { (\S. S = {}) } C { (\S. S = {}) }" +proof (rule consequence_rule) + let ?P = "general_union (\_. False)" + show "entails (\S. S = {}) ?P" + by (metis (full_types) Union_empty empty_iff entailsI general_unionI) + show "entails ?P (\S. S = {})" + using entailsI[of ?P "\S. S = {}"] general_union_def Sup_bot_conv(2) by metis + show "\ {general_union (\_. False)} C {general_union (\_. False)}" + proof (rule rule_BigUnion) + show "\ {(\_. False)} C {\_. False}" + using rule_False by auto + qed +qed + +definition has_subset where + "has_subset P S \ (\S'. S' \ S \ P S')" + +lemma has_subset_join_same: + "entails (has_subset P) (join P (\_. True))" + "entails (join P (\_. True)) (has_subset P)" + using entailsI[of "has_subset P" "join P (\_. True)"] has_subset_def join_def sup.absorb_iff2 + apply metis + using UnCI entails_def[of "join P (\_. True)" "has_subset P"] has_subset_def join_def subset_eq + by metis + + +(* derived from join *) +proposition rule_AtLeast: + assumes "\ {P} C {Q}" + shows "\ {has_subset P} C {has_subset Q}" +proof (rule consequence_rule) + let ?P = "join P (\_. True)" + let ?Q = "join Q (\_. True)" + show "\ {?P} C {?Q}" + by (simp add: assms rule_True rule_Union) +qed (auto simp add: has_subset_join_same) + +definition has_superset where + "has_superset P S \ (\S'. S \ S' \ P S')" + +proposition rule_AtMost: + assumes "\ {P} C {Q}" + shows "\ {has_superset P} C {has_superset Q}" +proof (rule hyper_hoare_tripleI) + fix S :: "(('a \ 'b) \ ('c \ 'd)) set" + assume "has_superset P S" + then obtain S' where "S \ S'" "P S'" + by (meson has_superset_def) + then have "Q (sem C S')" + using assms hyper_hoare_tripleE by blast + then show "has_superset Q (sem C S)" + by (meson \S \ S'\ has_superset_def sem_monotonic) +qed + + + + +subsection \Synchronous Reasoning (Proposition 14, Appendix H).\ + +theorem if_sync_rule: + assumes "\ {P} C1 {P1}" + and "\ {P} C2 {P2}" + and "\ {combine from_nat x P1 P2} C {combine from_nat x R1 R2}" + and "\ {R1} C1' {Q1}" + and "\ {R2} C2' {Q2}" + + and "not_free_var_hyper x P1" + and "not_free_var_hyper x P2" + and "from_nat 1 \ from_nat 2" + + and "not_free_var_hyper x R1" + and "not_free_var_hyper x R2" + + shows "\ {P} If (Seq C1 (Seq C C1')) (Seq C2 (Seq C C2')) {join Q1 Q2}" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "P S" + have r0: "sem (stmt.If (Seq C1 (Seq C C1')) (Seq C2 (Seq C C2'))) S + = sem C1' (sem C (sem C1 S)) \ sem C2' (sem C (sem C2 S))" + by (simp add: sem_if sem_seq) + moreover have "P1 (sem C1 S) \ P2 (sem C2 S)" + using asm0 assms(1) assms(2) hyper_hoare_tripleE by blast + + let ?S1 = "(modify_lvar_to x (from_nat 1)) ` (sem C1 S)" + let ?S2 = "(modify_lvar_to x (from_nat 2)) ` (sem C2 S)" + let ?f1 = "Set.filter (\\. fst \ x = from_nat 1)" + let ?f2 = "Set.filter (\\. fst \ x = from_nat 2)" + + have "P1 ?S1 \ P2 ?S2" + by (meson \P1 (sem C1 S) \ P2 (sem C2 S)\ assms(6) assms(7) not_free_var_hyper_def) + moreover have rr1: "Set.filter (\\. fst \ x = from_nat 1) (?S1 \ ?S2) = ?S1" + using injective_then_ok[of "from_nat 1" "from_nat 2" ?S1 x] + by (metis (no_types, lifting) assms(8)) + moreover have rr2: "Set.filter (\\. fst \ x = from_nat 2) (?S1 \ ?S2) = ?S2" + using injective_then_ok[of "from_nat 2" "from_nat 1" ?S2 x] + by (metis (no_types, lifting) assms(8) sup_commute) + ultimately have "combine from_nat x P1 P2 (?S1 \ ?S2)" + by (metis combineI) + then have "combine from_nat x R1 R2 (sem C (?S1 \ ?S2))" + using assms(3) hyper_hoare_tripleE by blast + moreover have "?f1 (sem C (?S1 \ ?S2)) = sem C ?S1" + using recover_after_sem[of "from_nat 1" "from_nat 2" ?S1 x ?S2] assms(8) rr1 rr2 + member_filter[of _ "\\. fst \ x = from_nat 1"] member_filter[of _ "\\. fst \ x = from_nat 2"] + by metis + then have "R1 (sem C ?S1)" + by (metis (mono_tags) calculation combine_def) + then have "R1 (sem C (sem C1 S))" + by (metis assms(9) not_free_var_hyper_def sem_of_modify_lvar) + moreover have "?f2 (sem C (?S1 \ ?S2)) = sem C ?S2" + using recover_after_sem[of "from_nat 2" "from_nat 1" ?S2 x ?S1] assms(8) rr1 rr2 sup_commute[of ] + member_filter[of _ "\\. fst \ x = from_nat 1" "?S1 \ ?S2"] member_filter[of _ "\\. fst \ x = from_nat 2" "?S1 \ ?S2"] + by metis + then have "R2 (sem C ?S2)" + by (metis (mono_tags) calculation(1) combine_def) + then have "R2 (sem C (sem C2 S))" + by (metis assms(10) not_free_var_hyper_def sem_of_modify_lvar) + + then show "join Q1 Q2 (sem (stmt.If (Seq C1 (Seq C C1')) (Seq C2 (Seq C C2'))) S)" + by (metis (full_types) r0 assms(4) assms(5) calculation(2) hyper_hoare_tripleE join_def) +qed + + +definition update_lvar_set where + "update_lvar_set u e S = { ((fst \')(u := e \'), snd \') |\'. \' \ S}" + +lemma equal_outside_set_helper: + "equal_outside_set {u} (fst \) (fst ((fst \)(u := x), snd \))" + by (simp add: equal_outside_set_def) + +lemma same_update_lvar_set: + "same_mod_updates {u} S (update_lvar_set u e S)" +proof (rule same_mod_updatesI) + show "\\'. \' \ update_lvar_set u e S \ \\\S. snd \ = snd \' \ equal_outside_set {u} (fst \') (fst \)" + proof - + fix \' assume "\' \ update_lvar_set u e S" + then obtain \' where "\' = ((fst \')(u := e \'), snd \')" "\' \ S" + using mem_Collect_eq update_lvar_set_def[of u e S] by auto + then show "\\\S. snd \ = snd \' \ equal_outside_set {u} (fst \') (fst \)" + by (meson equal_outside_set_helper equal_outside_sym snd_eqD) + qed + show "\\. \ \ S \ \\'\update_lvar_set u e S. snd \ = snd \' \ equal_outside_set {u} (fst \) (fst \')" + by (metis (mono_tags, lifting) equal_outside_set_helper mem_Collect_eq snd_conv update_lvar_set_def) +qed + + +lemma same_mod_updates_empty: + assumes "same_mod_updates vars {} S'" + shows "S' = {}" + by (meson assms equals0D equals0I same_mod_updates_def subset_mod_updatesE) + +definition not_fv_hyper where + "not_fv_hyper t A \ (\S S'. same_mod_updates {t} S S' \ A S \ A S')" + +lemma not_fv_hyperE: + assumes "not_fv_hyper e I" + and "same_mod_updates {e} S S'" + and "I S" + shows "I S'" + by (meson assms(1) assms(2) assms(3) not_fv_hyper_def) + +definition assign_exp_to_lvar where + "assign_exp_to_lvar e l \ = ((fst \)(l := e (snd \)), snd \)" + +definition assign_exp_to_lvar_set where + "assign_exp_to_lvar_set e l S = assign_exp_to_lvar e l ` S" + +lemma same_outside_set_lvar_assign_exp: + "snd \ = snd (assign_exp_to_lvar e l \) \ equal_outside_set {l} (fst \) (fst (assign_exp_to_lvar e l \))" + by (simp add: assign_exp_to_lvar_def equal_outside_set_def) + + +lemma assign_exp_to_lvar_set_same_mod_updates: + "same_mod_updates {l} S (assign_exp_to_lvar_set e l S)" +proof (rule same_mod_updatesI) + show "\\. \ \ S \ \\'\assign_exp_to_lvar_set e l S. snd \ = snd \' \ equal_outside_set {l} (fst \) (fst \')" + by (metis assign_exp_to_lvar_set_def rev_image_eqI same_outside_set_lvar_assign_exp) + show "\\'. \' \ assign_exp_to_lvar_set e l S \ \\\S. snd \ = snd \' \ equal_outside_set {l} (fst \') (fst \)" + by (metis (mono_tags, opaque_lifting) assign_exp_to_lvar_set_def equal_outside_sym imageE same_outside_set_lvar_assign_exp) +qed + + +lemma holds_forall_same_mod_updates: + assumes "same_mod_updates vars S S'" + and "holds_forall b S" + shows "holds_forall b S'" +proof (rule holds_forallI) + fix \' assume asm0: "\' \ S'" + then obtain \ where "\ \ S" "snd \ = snd \'" + by (metis assms(1) same_mod_updates_def subset_mod_updatesE) + then show "b (snd \')" + by (metis assms(2) holds_forall_def) +qed + +lemma not_fv_hyper_assign_exp: + assumes "not_fv_hyper t A" + shows "A S \ A (assign_exp_to_lvar_set e t S)" + by (metis assign_exp_to_lvar_set_same_mod_updates assms not_fv_hyper_def same_mod_updates_sym) + +lemma holds_forall_same_assign_lvar: + "holds_forall b S \ holds_forall b (assign_exp_to_lvar_set e l S)" (is "?A \ ?B") +proof + show "?A \ ?B" + by (simp add: assign_exp_to_lvar_def assign_exp_to_lvar_set_def holds_forall_def) + show "?B \ ?A" + by (simp add: assign_exp_to_lvar_def assign_exp_to_lvar_set_def holds_forall_def) +qed + +definition e_recorded_in_t where + "e_recorded_in_t e t S \ (\\\S. fst \ t = e (snd \))" + +lemma e_recorded_in_tI: + assumes "\\. \\S \ fst \ t = e (snd \)" + shows "e_recorded_in_t e t S" + by (simp add: assms e_recorded_in_t_def) + +definition e_smaller_than_t where + "e_smaller_than_t e t lt S \ (\\\S. lt (e (snd \)) (fst \ t))" + + +lemma low_expI: + assumes "\\ \'. \ \ S \ \' \ S \ (e (snd \) = e (snd \'))" + shows "low_exp e S" + using low_exp_def assms by blast + +lemma low_exp_forall_same_mod_updates: + assumes "same_mod_updates vars S S'" + and "low_exp b S" + shows "low_exp b S'" +proof (rule low_expI) + fix \' \'' assume asm0: "\' \ S' \ \'' \ S'" + then obtain \ where "\ \ S" "snd \ = snd \'" + by (metis assms(1) same_mod_updates_def subset_mod_updatesE) + then show "b (snd \') = b (snd \'')" + using asm0 assms(1) assms(2) low_exp_def[of b S] same_mod_updates_def[of vars S S'] subset_mod_updatesE + by metis +qed + +lemma e_recorded_in_t_if_assigned: + "e_recorded_in_t e t (assign_exp_to_lvar_set e t S)" + by (simp add: assign_exp_to_lvar_def assign_exp_to_lvar_set_def e_recorded_in_t_def) + +lemma low_exp_commute_assign_lvar: + "low_exp b (assign_exp_to_lvar_set e t S) \ low_exp b S" (is "?A \ ?B") +proof + assume ?A + then show ?B + by (meson assign_exp_to_lvar_set_same_mod_updates low_exp_forall_same_mod_updates same_mod_updates_sym) +next + assume ?B + then show ?A + by (meson assign_exp_to_lvar_set_same_mod_updates low_exp_forall_same_mod_updates) +qed + +end \ No newline at end of file diff --git a/thys/HyperHoareLogic/ExamplesCompositionality.thy b/thys/HyperHoareLogic/ExamplesCompositionality.thy new file mode 100644 --- /dev/null +++ b/thys/HyperHoareLogic/ExamplesCompositionality.thy @@ -0,0 +1,472 @@ +section \Examples\ + +text \In this file, we prove the correctness of the two compositionality +proofs presented in Appendix D.2.\ + +theory ExamplesCompositionality + imports Logic Compositionality +begin + +definition low where + "low l S \ (\\1 \2. \1 \ S \ \2 \ S \ snd \1 l = snd \2 l)" + +subsection \Examples using the core rules.\ + +definition GNI where + "GNI l h S \ (\\1 \2. \1 \ S \ \2 \ S + \ (\\ \ S. snd \ h = snd \1 h \ snd \ l = snd \2 l))" + +lemma GNI_I: + assumes "\\1 \2. \1 \ S \ \2 \ S + \ (\\ \ S. snd \ h = snd \1 h \ snd \ l = snd \2 l)" + shows "GNI l h S" + by (simp add: GNI_def assms) + +subsection \Examples using the compositionality rules\ + + + +definition has_minimum :: "'c \ ('d \ 'd \ bool) \ ('a, 'b, 'c, 'd) chyperassertion" where + "has_minimum x leq S \ (\\\S. \\'\S. leq (snd \ x) (snd \' x))" + +lemma has_minimumI: + assumes "\ \ S" + and "\\'. \' \ S \ leq (snd \ x) (snd \' x)" + shows "has_minimum x leq S" + by (metis assms(1) assms(2) has_minimum_def) + +definition is_monotonic where + "is_monotonic i x one two leq S \ (\\\S. \\'\S. fst \ i = one \ fst \' i = two \ leq (snd \ x) (snd \' x))" + +lemma is_monotonicI: + assumes "\\ \'. \ \ S \ \' \ S \ fst \ i = one \ fst \' i = two \ leq (snd \ x) (snd \' x)" + shows "is_monotonic i x one two leq S" + by (simp add: assms is_monotonic_def) + + + +lemma update_logical_equal_outside: + "equal_outside_set {i} (snd \) (snd (update_logical \ i v))" + by (simp add: equal_outside_set_def update_logical_def) + +lemma update_logical_read: + "fst (update_logical \ i v) i = v" + by (simp add: update_logical_def) + +lemma snd_update_logical_same: + "snd (update_logical \ i v) = snd \" + by (simp add: update_logical_def) + +text \Figure 12\ +proposition composing_monotonicity_and_minimum: + fixes P :: "((('a \ 'b) \ ('c \ 'd)) set \ bool)" + fixes i :: 'a + fixes x :: 'c + fixes y :: 'c + fixes leq :: "'d \ 'd \ bool" + fixes one :: 'b + fixes two :: 'b + + assumes "\ { P } C1 { has_minimum x leq }" + and "\ { is_monotonic i x one two leq } C2 { is_monotonic i y one two leq }" + and "\ { (is_singleton :: ((('a \ 'b) \ ('c \ 'd)) set \ bool)) } C2 { is_singleton }" + and "one \ two" + + and "\x. leq x x" \\reflexivity\ + + shows "\ { P } C1 ;; C2 { has_minimum y leq }" + using assms(1) +proof (rule seq_rule) + + let ?P1 = "is_singleton \ (Set.filter (\\. fst \ i = one))" + let ?P2 = "is_monotonic i x one two leq" + let ?P3 = "\S. \\\S. fst \ i = one \ fst \ i = two" + + let ?P = "conj ?P1 (conj ?P2 ?P3)" + + let ?Q1 = "is_singleton \ (Set.filter (\\. fst \ i = one))" + let ?Q2 = "is_monotonic i y one two leq" + let ?Q3 = "\S. \\\S. fst \ i = one \ fst \ i = two" + + let ?Q = "conj ?Q1 (conj ?Q2 ?Q3)" + + + show "\ { (has_minimum x leq :: ((('a \ 'b) \ ('c \ 'd)) set \ bool)) } C2 { has_minimum y leq }" + proof (rule rule_LUpdate) + + show "entails_with_updates {i} (has_minimum x leq) ?P" + proof (rule entails_with_updatesI) + fix S :: "(('a \ 'b) \ ('c \ 'd)) set" + assume asm0: "has_minimum x leq S" + then obtain \ where minimum: "\ \ S" "\\'. \' \ S \ leq (snd \ x) (snd \' x)" + by (metis has_minimum_def) + let ?\ = "update_logical \ i one" + let ?S = "{ update_logical \' i two |\'. \' \ S } \ {?\}" + have "same_mod_updates {i} S ?S" + proof (rule same_mod_updatesI) + fix \' assume asm1: "\' \ S" + then have "update_logical \' i two \ ?S" + by blast + moreover have "snd \' = snd (update_logical \' i two) \ equal_outside_set {i} (fst \') (fst (update_logical \' i two))" + by (metis equal_outside_setI fst_eqD fun_upd_other singletonI snd_update_logical_same update_logical_def) + ultimately show "\\''\?S. snd \' = snd \'' \ equal_outside_set {i} (fst \') (fst \'')" + by blast + next + fix \' + assume asm1: "\' \ {update_logical \' i two |\'. \' \ S} \ {update_logical \ i one}" + show "\\\S. snd \ = snd \' \ equal_outside_set {i} (fst \') (fst \)" + proof (cases "\' \ {update_logical \' i two |\'. \' \ S}") + case True + then obtain \'' where "\' = update_logical \'' i two" "\'' \ S" + by blast + then show ?thesis + by (metis (mono_tags, lifting) equal_outside_set_def fst_conv fun_upd_other insertCI snd_update_logical_same update_logical_def) + next + case False + then show ?thesis + by (metis (mono_tags, lifting) UnE asm1 equal_outside_setI fst_eqD fun_upd_other minimum(1) singletonD singletonI snd_update_logical_same update_logical_def) + qed + qed + moreover have "is_singleton (Set.filter (\\. fst \ i = one) ?S)" + proof - + have "Set.filter (\\. fst \ i = one) ?S \ {?\}" + proof + fix \ assume "\ \ Set.filter (\\. fst \ i = one) ?S" + then have "\ \ ?S \ fst \ i = one" + by simp + then show "\ \ {?\}" + using assms(4) update_logical_read by force + qed + moreover have "{?\} \ Set.filter (\\. fst \ i = one) ?S" + by (simp add: update_logical_read) + ultimately show ?thesis + by (simp add: is_singleton_def order_antisym_conv) + qed + moreover have "is_monotonic i x one two leq ?S" + proof (rule is_monotonicI) + fix \' \'' assume asm1: "\' \ ?S" "\'' \ ?S" "fst \' i = one \ fst \'' i = two" + then have " \' \ {update_logical \ i one} \ \'' \ {update_logical \' i two |\'. \' \ S}" + using assms(4) update_logical_read by fastforce + then show "leq (snd \' x) (snd \'' x)" + by (metis (no_types, lifting) asm1(2) calculation(1) minimum(2) same_mod_updates_def singletonD snd_update_logical_same subset_mod_updatesE) + qed + moreover have "\\'. \' \ ?S \ fst \' i = one \ fst \' i = two" + using update_logical_read by fastforce + ultimately have "same_mod_updates {i} S ?S \ ?P ?S" + using comp_eq_dest_lhs[of is_singleton "Set.filter (\\. fst \ i = one)"] conj_def[of ?P1 "conj ?P2 ?P3"] conj_def[of ?P2 ?P3] + by simp + then show "\S'. same_mod_updates {i} S S' \ ?P S'" + by blast + qed + + show "invariant_on_updates {i} (has_minimum y leq :: ((('a \ 'b) \ ('c \ 'd)) set \ bool))" + proof (rule invariant_on_updatesI) + fix S :: "(('a \ 'b) \ ('c \ 'd)) set" + fix S' + assume asm0: "same_mod_updates {i} S S'" "has_minimum y leq S" + then show "has_minimum y leq S'" + using has_minimum_def[of y leq S'] has_minimum_def[of y leq S] same_mod_updates_def[of "{i}" S S'] subset_mod_updatesE[of "{i}"] + by metis + qed + + show "\ { ?P } C2 { has_minimum y leq }" + proof (rule consequence_rule) + show "entails ?Q (has_minimum y leq)" + proof (rule entailsI) + fix S :: "(('a \ 'b) \ ('c \ 'd)) set" + assume "?Q S" + then obtain asm0: "is_singleton (Set.filter (\\. fst \ i = one) S)" "is_monotonic i y one two leq S" + "\\'. \' \ S \ fst \' i = one \ fst \' i = two" + by (simp add: conj_def) + then obtain \ where "Set.filter (\\. fst \ i = one) S = {\}" + by (metis is_singleton_def) + then have "\ \ S" by auto + then show "has_minimum y leq S" + proof (rule has_minimumI) + fix \' + assume asm1: "\' \ S" + show "leq (snd \ y) (snd \' y)" + proof (cases "fst \' i = one") + case True + then show ?thesis + using \Set.filter (\\. fst \ i = one) S = {\}\ asm1 assms(5) by fastforce + next + case False + then have "fst \' i = two" + using asm0(3) asm1 by blast + then show ?thesis + using \Set.filter (\\. fst \ i = one) S = {\}\ asm0(2) asm1 is_monotonic_def[of i y one two leq S] member_filter singleton_iff + by force + qed + qed + qed + show "\ { ?P } C2 { ?Q }" + proof (rule rule_and3) + show " \ {is_singleton \ Set.filter (\\. fst \ i = one)} C2 {is_singleton \ Set.filter (\\. fst \ i = one)}" + proof (rule rule_apply) + show "\ { (is_singleton :: ((('a \ 'b) \ ('c \ 'd)) set \ bool)) } C2 {is_singleton}" + using assms(3) by blast + show "commute_with_sem (Set.filter (\\. fst \ i = one))" + by (simp add: filter_prop_commute) + qed + show "\ {is_monotonic i x one two leq} C2 {is_monotonic i y one two leq}" + by (simp add: assms(2)) + + show "\ {(\S. \\\S. fst \ i = one \ fst \ i = two)} C2 {\S. \\\S. fst \ i = one \ fst \ i = two}" + using rule_lframe_single by fast + qed + qed (auto simp add: entails_refl) + qed +qed + +text \In this definition, we use a logical variable for h, which records the initial value of the program variable h\ + +definition lGNI :: "'pvar \ 'lvar \ (('lvar, 'lval, 'pvar, 'pval) state) set \ bool" where + "lGNI l h S \ (\\1 \ S. (\\2 \ S. \\ \ S. fst \ h = fst \1 h \ snd \ l = snd \2 l))" + +text \Figure 13\ +proposition composing_GNI_with_SNI: + fixes h :: 'lvar + fixes l :: 'pvar + + assumes "\ { (low l :: (('lvar, 'lval, 'pvar, 'pval) state) hyperassertion) } C2 { low l }" + and "\ { (not_empty :: (('lvar, 'lval, 'pvar, 'pval) state) hyperassertion) } C2 { not_empty }" + and "\ { (low l :: (('lvar, 'lval, 'pvar, 'pval) state) hyperassertion) } C1 { lGNI l h }" + shows "\ { (low l :: (('lvar, 'lval, 'pvar, 'pval) state) hyperassertion) } C1;; C2 { lGNI l h }" + using assms(3) +proof (rule seq_rule) + show "\ { (lGNI l h :: (('lvar, 'lval, 'pvar, 'pval) state) hyperassertion)} C2 {lGNI l h}" + unfolding lGNI_def + proof (rule rule_linking) + fix \1 \1' :: "('lvar, 'lval, 'pvar, 'pval) state" + assume asm0: "fst \1 = fst \1' \ \ { (in_set \1 :: ('lvar, 'lval, 'pvar, 'pval) state hyperassertion)} C2 {in_set \1'}" + show "\ {(\S. \\2\S. \\\S. fst \ h = fst \1 h \ snd \ l = snd \2 l)} C2 {\S. \\2\S. \\\S. fst \ h = fst \1' h \ snd \ l = snd \2 l}" + proof (rule consequence_rule) + let ?ex = "\S. \\ \ S. fst \ h = fst \1 h" + let ?P = "general_union (conj (low l) ?ex)" + show "\ { (?P :: (('lvar, 'lval, 'pvar, 'pval) state) hyperassertion) } C2 {?P}" + proof (rule rule_BigUnion; rule rule_And) + show "\ {(low l :: (('lvar, 'lval, 'pvar, 'pval) state) hyperassertion)} C2 {low l}" + using assms(1) by blast + show "\ { ?ex } C2 { ?ex }" + proof (rule consequence_rule) + let ?b = "\\0. \0 h = fst \1 h" + show "\ {not_empty \ Set.filter (?b \ fst)} C2 { not_empty \ Set.filter (?b \ fst)}" + using assms(2) rule_LFilter by auto + show "entails (\S. \\\S. fst \ h = fst \1 h) (not_empty \ Set.filter ((\\0. \0 h = fst \1 h) \ fst))" + using CollectI Set.filter_def comp_apply empty_iff not_empty_def + entailsI[of "(\S. \\\S. fst \ h = fst \1 h)" "(not_empty \ Set.filter ((\\0. \0 h = fst \1 h) \ fst))"] + by fastforce + show "entails (not_empty \ Set.filter ((\\0. \0 h = fst \1 h) \ fst)) (\S. \\\S. fst \ h = fst \1 h)" + proof (rule entailsI) + fix S assume "(not_empty \ Set.filter ((\\0. \0 h = fst \1 h) \ fst)) S" + then obtain \ where "\ \ Set.filter ((\\0. \0 h = fst \1 h) \ fst) S" + by (metis comp_apply equals0I not_empty_def) + then show "\\\S. fst \ h = fst \1 h" + by auto + qed + qed + qed + show "entails (\S. \\2\S. \\\S. fst \ h = fst \1 h \ snd \ l = snd \2 l) ?P" + proof (rule entailsI) + fix S assume asm0: "\\2\S. \\\S. fst \ h = fst \1 h \ snd \ l = snd \2 l" + thm general_unionI + let ?F = "{ {\, \2} |\ \2. \ \ S \ \2 \ S \ snd \ l = snd \2 l \ fst \ h = fst \1 h }" + show "general_union (Logic.conj (low l) (\S. \\\S. fst \ h = fst \1 h)) S" + proof (rule general_unionI) + show "S = \ ?F" + proof + show "S \ \ ?F" + proof + fix \2 assume "\2 \ S" + then obtain \ where "\\S" "fst \ h = fst \1 h \ snd \ l = snd \2 l" + using asm0 by blast + then have "{\, \2} \ ?F" + using \\2 \ S\ by blast + then show "\2 \ \ ?F" by blast + qed + show "\ ?F \ S" by blast + qed + fix S' assume asm1: "S' \ {{\, \2} |\ \2. \ \ S \ \2 \ S \ snd \ l = snd \2 l \ fst \ h = fst \1 h}" + then obtain \ \2 where "S' = {\, \2}" "\ \ S \ \2 \ S \ snd \ l = snd \2 l \ fst \ h = fst \1 h" + by blast + then show "Logic.conj (low l) (\S. \\\S. fst \ h = fst \1 h) S'" + using conj_def[of "low l" "\S. \\\S. fst \ h = fst \1 h"] insert_iff low_def singletonD + by fastforce + qed + qed + show "entails ?P (\S. \\2\S. \\\S. fst \ h = fst \1' h \ snd \ l = snd \2 l)" + proof (rule entailsI) + fix S assume "general_union (Logic.conj (low l) (\S. \\\S. fst \ h = fst \1 h)) S" + then obtain F where "S = \ F" "\S'. S' \ F \ low l S' \ (\\\S'. fst \ h = fst \1 h)" + using conj_def[of "low l" "\S. \\\S. fst \ h = fst \1 h"] + general_unionE[of "Logic.conj (low l) (\S. \\\S. fst \ h = fst \1 h)" S] + by (metis (mono_tags, lifting)) + then show "\\2\S. \\\S. fst \ h = fst \1' h \ snd \ l = snd \2 l" + by (metis (mono_tags, lifting) Union_iff asm0 low_def) + qed + qed + qed +qed + + +subsection \Other examples\ + +lemma program_1_sat_gni: + assumes "y \ l \ y \ h \ l \ h" + shows "\ { low l } Seq (Havoc y) (Assign l (\\. (\ h :: int) + \ y)) { GNI l h }" +proof (rule RuleSeq) + let ?R = "\S. \\1 \2. \1 \ S \ \2 \ S + \ (\\ \ S. (snd \ h :: int) = snd \1 h \ snd \ h + snd \ y = snd \2 h + snd \2 y )" + + show "\ { low l } Havoc y {?R}" + proof (rule RuleCons) + show "\ {(\S. ?R {(l, \(y := v)) |l \ v. (l, \) \ S})} Havoc y {?R}" + using RuleHavoc[of ?R] by blast + show "entails (low l) (\S. ?R {(l, \(y := v)) |l \ (v :: int). (l, \) \ S})" + proof (rule entailsI) + fix S + show "?R {(l, \(y := v)) |l \ (v :: int). (l, \) \ S}" + proof (clarify) + fix a b aa ba l la \ \' v va + assume asm0: "(l, \) \ S" "(la, \') \ S" + let ?v = "(\'(y := va)) h + (\'(y := va)) y + - \ h" + let ?\ = "(l, \(y := ?v))" + have "snd ?\ h = snd (l, \(y := v)) h \ snd ?\ h + snd ?\ y = snd (la, \'(y := va)) h + snd (la, \'(y := va)) y" + using assms by force + then show "\\\{(l, \(y := v)) |l \ v. (l, \) \ S}. + snd \ h = snd (l, \(y := v)) h \ snd \ h + snd \ y = snd (la, \'(y := va)) h + snd (la, \'(y := va)) y" + using asm0(1) by blast + qed + qed + show "entails ?R ?R" + by (meson entailsI) + qed + show "\ {?R} (Assign l (\\. \ h + \ y)) {GNI l h}" + proof (rule RuleCons) + show "\ {(\S. GNI l h {(la, \(l := \ h + \ y)) |la \. (la, \) \ S})} Assign l (\\. \ h + \ y) {GNI l h}" + using RuleAssign[of "GNI l h" l "\\. \ h + \ y"] by blast + show "entails (GNI l h) (GNI l h)" + by (simp add: entails_def) + show "entails ?R (\S. GNI l h {(la, \(l := \ h + \ y)) |la \. (la, \) \ S})" + proof (rule entailsI) + fix S + assume asm0: "\\1 \2. \1 \ S \ \2 \ S \ (\\\S. snd \ h = snd \1 h \ snd \ h + snd \ y = snd \2 h + snd \2 y)" + show "GNI l h {(la, \(l := \ h + \ y)) |la \. (la, \) \ S}" + proof (rule GNI_I) + fix \1 \2 + assume asm1: "\1 \ {(la, \(l := \ h + \ y)) |la \. (la, \) \ S} \ \2 \ {(la, \(l := \ h + \ y)) |la \. (la, \) \ S}" + then obtain la \ la' \' where "(la, \) \ S" "(la', \') \ S" "\1 = (la, \(l := \ h + \ y))" "\2 = (la', \'(l := \' h + \' y))" + by blast + then obtain \ where "\\S" "snd \ h = \ h" "snd \ h + snd \ y = \' h + \' y" + using asm0 snd_conv by force + let ?\ = "(fst \, (snd \)(l := snd \ h + snd \ y))" + have "snd ?\ h = snd \1 h \ snd ?\ l = snd \2 l" + using \\1 = (la, \(l := \ h + \ y))\ \\2 = (la', \'(l := \' h + \' y))\ + \snd \ h + snd \ y = \' h + \' y\ \snd \ h = \ h\ assms by force + then show "\\\{(la, \(l := \ h + \ y)) |la \. (la, \) \ S}. snd \ h = snd \1 h \ snd \ l = snd \2 l" + using \\ \ S\ mem_Collect_eq[of ?\] + by (metis (mono_tags, lifting) prod.collapse) + qed + qed + qed +qed + + +lemma program_2_violates_gni: + assumes "y \ l \ y \ h \ l \ h" + shows "\ { conj (low l) (\S. \a \ S. \b \ S. (snd a h :: nat) \ snd b h) } + Seq (Seq (Havoc y) (Assume (\\. \ y \ (0 :: nat) \ \ y \ (100 :: nat)))) (Assign l (\\. \ h + \ y)) + {\(S :: (('lvar \ 'lval) \ ('a \ nat)) set). \ GNI l h S}" +proof (rule RuleSeq) + + let ?R0 = "\(S :: (('lvar \ 'lval) \ ('a \ nat)) set). + (\a \ S. \b \ S. snd b h > snd a h \ snd a y \ (0 :: nat) \ snd a y \ 100 \ snd b y = 100)" + let ?R1 = "\(S :: (('lvar \ 'lval) \ ('a \ nat)) set). + (\a \ S. \b \ S. snd b h > snd a h \ snd b y = 100) \ (\c \ S. snd c y \ 100)" + let ?R2 = "\(S :: (('lvar \ 'lval) \ ('a \ nat)) set). + \a \ S. \b \ S. \c \ S. snd c h = snd a h \ snd c h + snd c y = snd b h + snd b y" + + show "\ { conj (low l) (\S. \a\S. \b\S. snd a h \ snd b h)} Seq (Havoc y) (Assume (\\. 0 \ \ y \ \ y \ (100 :: nat))) {?R1}" + proof (rule RuleSeq) + show "\ {conj (low l) (\S. \a\S. \b\S. snd a h \ snd b h)} Havoc y { ?R0 }" + proof (rule RuleCons) + show "\ {(\S. ?R0 {(l, \(y := v)) |l \ v. (l, \) \ S})} Havoc y {?R0}" + using RuleHavoc[of _ y] by fast + show "entails ?R0 ?R0" + by (simp add: entailsI) + show "entails (conj (low l) (\S. \a\S. \b\S. snd a h \ snd b h)) (\S. ?R0 {(l, \(y := v)) |l \ v. (l, \) \ S})" + proof (rule entailsI) + fix S :: "(('lvar \ 'lval) \ ('a \ nat)) set" + assume "conj (low l) (\S. \a\S. \b\S. snd a h \ snd b h) S" + then have "\a\S. \b\S. snd a h \ snd b h" using conj_def[of "low l" "\S. \a\S. \b\S. snd a h \ snd b h"] + by blast + then obtain a b where "a \ S" "b \ S" "snd b h > snd a h" + by (meson linorder_neq_iff) + let ?a = "(fst a, (snd a)(y := 100))" + let ?b = "(fst b, (snd b)(y := 100))" + have "?a \ {(l, \(y := v)) |l \ v. (l, \) \ S} \ ?b \ {(l, \(y := v)) |l \ v. (l, \) \ S}" + using \a \ S\ \b \ S\ by fastforce + moreover have "snd ?b h > snd ?a h \ snd ?a y \ (0 :: nat) \ snd ?a y \ 100 \ snd ?b y = 100" + using \snd a h < snd b h\ assms by force + ultimately show "?R0 {(l, \(y := v)) |l \ v. (l, \) \ S}" by blast + qed + qed + show "\ {?R0} Assume (\\. 0 \ \ y \ \ y \ 100) {?R1}" + proof (rule RuleCons) + show "\ {(\S. ?R1 (Set.filter ((\\. 0 \ \ y \ \ y \ 100) \ snd) + S))} Assume (\\. 0 \ \ y \ \ y \ 100) {?R1}" + using RuleAssume[of _ "\\. 0 \ \ y \ \ y \ 100"] + by fast + show "entails ?R1 ?R1" + by (simp add: entailsI) + show "entails ?R0 (\S. ?R1 (Set.filter ((\\. 0 \ \ y \ \ y \ 100) \ snd) S))" + proof (rule entailsI) + fix S :: "(('lvar \ 'lval) \ ('a \ nat)) set" + assume asm0: "?R0 S" + then obtain a b where "a\S" "b\S" "snd a h < snd b h \ 0 \ snd a y \ snd a y \ (100 :: nat) \ snd b y = 100 " + by blast + then have "a \ Set.filter ((\\. 0 \ \ y \ \ y \ 100) \ snd) S \ b \ Set.filter ((\\. 0 \ \ y \ \ y \ 100) \ snd) S" + by (simp add: \a \ S\ \b \ S\) + then show "?R1 (Set.filter ((\\. 0 \ \ y \ \ y \ 100) \ snd) S)" + using \snd a h < snd b h \ 0 \ snd a y \ snd a y \ 100 \ snd b y = 100\ by force + qed + qed + qed + show "\ { ?R1 } Assign l (\\. \ h + \ y) {\S. \ GNI l h S}" + proof (rule RuleCons) + show "\ {(\S. \ GNI l h {(la, \(l := \ h + \ y)) |la \. (la, \) \ S})} Assign l (\\. \ h + \ y) {\S. \ GNI l h S}" + using RuleAssign[of "\S. \ GNI l h S" l "\\. \ h + \ y"] + by blast + show "entails (\S. \ GNI l h S) (\S. \ GNI l h S)" + by (simp add: entails_def) + show "entails (\S. (\a\S. \b\S. snd a h < snd b h \ snd b y = 100) \ (\c\S. snd c y \ 100)) + (\(S :: (('lvar \ 'lval) \ ('a \ nat)) set). \ GNI l h {(la, \(l := \ h + \ y)) |la \. (la, \) \ S})" + proof (rule entailsI) + fix S :: "(('lvar \ 'lval) \ ('a \ nat)) set" + assume asm0: "(\a\S. \b\S. snd a h < snd b h \ snd b y = 100) \ (\c\S. snd c y \ 100)" + then obtain a b where asm1: "a\S" "b\S" "snd a h < snd b h \ snd b y = 100" by blast + let ?a = "(fst a, (snd a)(l := snd a h + snd a y))" + let ?b = "(fst b, (snd b)(l := snd b h + snd b y))" + have "\la \. (la, \) \ S \ (\(l := \ h + \ y)) h = snd ?a h \ (\(l := \ h + \ y)) l \ snd ?b l" + using asm0 asm1(3) assms by fastforce + moreover have r: "?a \ {(la, \(l := \ h + \ y)) |la \. (la, \) \ S} \ ?b \ {(la, \(l := \ h + \ y)) |la \. (la, \) \ S}" + using asm1(1) asm1(2) by fastforce + show "\ GNI l h {(la, \(l := \ h + \ y)) |la \. (la, \) \ S}" + proof (rule ccontr) + assume "\ \ GNI l h {(la, \(l := \ h + \ y)) |la \. (la, \) \ S}" + then have "GNI l h {(la, \(l := \ h + \ y)) |la \. (la, \) \ S}" + by blast + then obtain \ where "\ \ {(la, \(l := \ h + \ y)) |la \. (la, \) \ S}" "snd \ h = snd ?a h" "snd \ l = snd ?b l" + using GNI_def[of l h "{(la, \(l := \ h + \ y)) |la \. (la, \) \ S}"] r + by meson + then show False + using calculation by auto + qed + qed + qed +qed + + + +end \ No newline at end of file diff --git a/thys/HyperHoareLogic/Expressivity.thy b/thys/HyperHoareLogic/Expressivity.thy --- a/thys/HyperHoareLogic/Expressivity.thy +++ b/thys/HyperHoareLogic/Expressivity.thy @@ -1,1871 +1,1943 @@ -text \In this file, we prove most results of section V: hyper-triples subsume many other triples, -as well as example 4.\ +text \In this file, we prove most results of Appendix C: + hyper-triples subsume many other triples, as well as example 3.\ theory Expressivity imports ProgramHyperproperties begin subsection \Hoare Logic (HL)~\cite{HoareLogic}\ -text \Definition 8\ +paragraph \Definition 16\ definition HL where "HL P C Q \ (\\ \' l. (l, \) \ P \ (\C, \\ \ \') \ (l, \') \ Q)" lemma HLI: assumes "\\ \' l. (l, \) \ P \ \C, \\ \ \' \ (l, \') \ Q" shows "HL P C Q" using assms HL_def by blast lemma hoarifyI: assumes "\\ \'. (\, \') \ S \ \ \ P \ \' \ Q" shows "hoarify P Q S" by (metis assms hoarify_def prod.collapse) definition HL_hyperprop where "HL_hyperprop P Q S \ (\l. \p \ S. (l, fst p) \ P \ (l, snd p) \ Q)" lemma connection_HL: "HL P C Q \ HL_hyperprop P Q (set_of_traces C)" (is "?A \ ?B") proof assume ?A then show ?B by (simp add: HL_def HL_hyperprop_def set_of_traces_def) next assume ?B show ?A proof (rule HLI) fix \ \' l assume asm0: "(l, \) \ P" "\C, \\ \ \'" then have "(\, \') \ set_of_traces C" by (simp add: set_of_traces_def) then show "(l, \') \ Q" using \HL_hyperprop P Q (set_of_traces C)\ asm0(1) HL_hyperprop_def by fastforce qed qed -text \Proposition 1\ +paragraph \Proposition 1\ theorem HL_expresses_hyperproperties: "\H. (\C. hypersat C H \ HL P C Q) \ k_hypersafety 1 H" proof - let ?H = "HL_hyperprop P Q" have "\C. hypersat C ?H \ HL P C Q" by (simp add: connection_HL hypersat_def) moreover have "k_hypersafety 1 ?H" proof (rule k_hypersafetyI) fix S assume asm0: "\ HL_hyperprop P Q S" then obtain l p where "p \ S" "(l, fst p) \ P" "(l, snd p) \ Q" using HL_hyperprop_def by blast let ?S = "{p}" have "max_k 1 ?S \ (\S''. ?S \ S'' \ \ HL_hyperprop P Q S'')" by (metis (no_types, lifting) One_nat_def \(l, fst p) \ P\ \(l, snd p) \ Q\ card.empty card.insert empty_iff finite.intros(1) finite.intros(2) le_numeral_extra(4) max_k_def HL_hyperprop_def singletonI subsetD) then show "\S'\S. max_k 1 S' \ (\S''. S' \ S'' \ \ HL_hyperprop P Q S'')" by (meson \p \ S\ empty_subsetI insert_subsetI) qed ultimately show ?thesis by blast qed -text \Proposition 2\ +paragraph \Proposition 2\ theorem encoding_HL: "HL P C Q \ (hyper_hoare_triple (over_approx P) C (over_approx Q))" (is "?A \ ?B") proof (rule iffI) show "?B \ ?A" proof - assume asm0: ?B show ?A proof (rule HLI) fix \ \' l assume asm1: "(l, \) \ P" "\C, \\ \ \'" then have "over_approx P {(l, \)}" by (simp add: over_approx_def) then have "(over_approx Q) (sem C {(l, \)})" using asm0 hyper_hoare_tripleE by auto then show "(l, \') \ Q" by (simp add: asm1(2) in_mono in_sem over_approx_def) qed qed next assume r: ?A show ?B proof (rule hyper_hoare_tripleI) fix S assume asm0: "over_approx P S" then have "S \ P" by (simp add: over_approx_def) then have "sem C S \ sem C P" by (simp add: sem_monotonic) then have "sem C S \ Q" using r HL_def[of P C Q] by (metis (no_types, lifting) fst_conv in_mono in_sem snd_conv subrelI) then show "over_approx Q (sem C S)" by (simp add: over_approx_def) qed qed lemma entailment_order_hoare: assumes "P \ P'" shows "entails (over_approx P) (over_approx P')" by (simp add: assms entails_def over_approx_def subset_trans) subsection \Cartesian Hoare Logic (CHL)~\cite{CHL16}\ -text \Notation 3\ definition k_sem where "k_sem C states states' \ (\i. (fst (states i) = fst (states' i) \ single_sem C (snd (states i)) (snd (states' i))))" lemma k_semI: assumes "\i. (fst (states i) = fst (states' i) \ single_sem C (snd (states i)) (snd (states' i)))" shows "k_sem C states states'" by (simp add: assms k_sem_def) lemma k_semE: assumes "k_sem C states states'" shows "fst (states i) = fst (states' i) \ single_sem C (snd (states i)) (snd (states' i))" using assms k_sem_def by fastforce -text \Definition 9\ +paragraph \Definition 17\ definition CHL where "CHL P C Q \ (\states. states \ P \ (\states'. k_sem C states states' \ states' \ Q))" lemma CHLI: assumes "\states states'. states \ P \ k_sem C states states' \ states' \ Q" shows "CHL P C Q" by (simp add: assms CHL_def) lemma CHLE: assumes "CHL P C Q" and "states \ P" and "k_sem C states states'" shows "states' \ Q" using assms(1) assms(2) assms(3) CHL_def by fast definition encode_CHL where "encode_CHL from_nat x P S \ (\states. (\i. states i \ S \ fst (states i) x = from_nat i) \ states \ P)" lemma encode_CHLI: assumes "\states. (\i. states i \ S \ fst (states i) x = from_nat i) \ states \ P" shows "encode_CHL from_nat x P S" using assms(1) encode_CHL_def by force lemma encode_CHLE: assumes "encode_CHL from_nat x P S" and "\i. states i \ S" and "\i. fst (states i) x = from_nat i" shows "states \ P" by (metis assms(1) assms(2) assms(3) encode_CHL_def) lemma equal_change_lvar: assumes "fst \ x = y" shows "\ = ((fst \)(x := y), snd \)" using assms by fastforce -text \Proposition 3\ +paragraph \Proposition 3\ theorem encoding_CHL: assumes "not_free_var_of P x" and "not_free_var_of Q x" and "injective from_nat" shows "CHL P C Q \ \ {encode_CHL from_nat x P} C {encode_CHL from_nat x Q}" (is "?A \ ?B") proof assume ?A show ?B proof (rule hyper_hoare_tripleI) fix S assume "encode_CHL from_nat x P S" then obtain asm0: "\states states'. (\i. states i \ S) \ (\i. fst (states i) x = from_nat i) \ states \ P" by (simp add: encode_CHLE) show "encode_CHL from_nat x Q (sem C S)" proof (rule encode_CHLI) fix states' assume asm1: "\i. states' i \ sem C S \ fst (states' i) x = from_nat i" let ?states = "\i. (fst (states' i), SOME \. (fst (states' i), \) \ S \ single_sem C \ (snd (states' i)))" show "states' \ Q" using \?A\ proof (rule CHLE) show "?states \ P" proof (rule asm0) fix i let ?\ = "SOME \. ((fst (states' i), \) \ S \ \C, \\ \ snd (states' i))" have r: "(fst (states' i), ?\) \ S \ \C, ?\\ \ snd (states' i)" using someI_ex[of "\\. (fst (states' i), \) \ S \ \C, \\ \ snd (states' i)"] asm1 in_sem by blast then show "?states i \ S" by blast show "fst (?states i) x = from_nat i" by (simp add: asm1) qed show "k_sem C ?states states'" proof (rule k_semI) fix i let ?\ = "SOME \. ((fst (states' i), \) \ S \ \C, \\ \ snd (states' i))" have r: "(fst (states' i), ?\) \ S \ \C, ?\\ \ snd (states' i)" using someI_ex[of "\\. (fst (states' i), \) \ S \ \C, \\ \ snd (states' i)"] asm1 in_sem by blast then show "fst (?states i) = fst (states' i) \ \C, snd (?states i)\ \ snd (states' i)" by simp qed qed qed qed next assume asm0: "\ {encode_CHL from_nat x P} C {encode_CHL from_nat x Q}" show "CHL P C Q" proof (rule CHLI) fix states states' assume asm1: "states \ P" "k_sem C states states'" let ?states = "\i. ((fst (states i))(x := from_nat i), snd (states i))" let ?states' = "\i. ((fst (states i))(x := from_nat i), snd (states' i))" let ?S = "range ?states" have "encode_CHL from_nat x Q (sem C ?S)" using asm0 proof (rule hyper_hoare_tripleE) show "encode_CHL from_nat x P ?S" proof (rule encode_CHLI) fix f assume asm2: "\i. f i \ ?S \ fst (f i) x = from_nat i" have "f = ?states" proof (rule ext) fix i obtain j where j_def: "f i = ((fst (states j))(x := from_nat j), snd (states j))" using asm2 by fastforce then have "from_nat j = from_nat i" by (metis asm2 fst_conv fun_upd_same) then show "f i = ((fst (states i))(x := from_nat i), snd (states i))" by (metis j_def assms(3) injective_def) qed moreover have "?states \ P" using assms(1) proof (rule not_free_var_ofE) show "states \ P" using asm1(1) by simp fix i show "differ_only_by (fst (states i)) (fst ((fst (states i))(x := from_nat i), snd (states i))) x" by (simp add: differ_only_by_def) show "snd (states i) = snd ((fst (states i))(x := from_nat i), snd (states i))" by simp qed ultimately show "f \ P" by meson qed qed then have "?states' \ Q" proof (rule encode_CHLE) fix i show "fst ((fst (states i))(x := from_nat i), snd (states' i)) x = from_nat i" by simp moreover have "single_sem C (snd (?states i)) (snd (?states' i))" using asm1(2) k_sem_def by fastforce ultimately show "((fst (states i))(x := from_nat i), snd (states' i)) \ sem C ?S" using in_sem by fastforce qed show "states' \ Q" using assms(2) proof (rule not_free_var_ofE[of Q x]) show "?states' \ Q" by (simp add: \(\i. ((fst (states i))(x := from_nat i), snd (states' i))) \ Q\) fix i show "differ_only_by (fst ((fst (states i))(x := from_nat i), snd (states' i))) (fst (states' i)) x" by (metis asm1(2) diff_by_update fst_conv k_sem_def) qed (auto) qed qed definition CHL_hyperprop where "CHL_hyperprop P Q S \ (\l p. (\i. p i \ S) \ (\i. (l i, fst (p i))) \ P \ (\i. (l i, snd (p i))) \ Q)" lemma CHL_hyperpropI: assumes "\l p. (\i. p i \ S) \ (\i. (l i, fst (p i))) \ P \ (\i. (l i, snd (p i))) \ Q" shows "CHL_hyperprop P Q S" by (simp add: assms CHL_hyperprop_def) lemma CHL_hyperpropE: assumes "CHL_hyperprop P Q S" and "\i. p i \ S" and "(\i. (l i, fst (p i))) \ P" shows "(\i. (l i, snd (p i))) \ Q" using assms(1) assms(2) assms(3) CHL_hyperprop_def by blast -text \Proposition 10\ +paragraph \Proposition 3\ theorem CHL_hyperproperty: "hypersat C (CHL_hyperprop P Q) \ CHL P C Q" (is "?A \ ?B") proof assume ?A show ?B proof (rule CHLI) fix states states' assume asm0: "states \ P" "k_sem C states states'" let ?p = "\i. (snd (states i), snd (states' i))" let ?l = "\i. fst (states i)" have "range ?p \ set_of_traces C" proof (rule subsetI) fix x assume "x \ range ?p" then obtain i where "x = (snd (states i), snd (states' i))" by blast then show "x \ set_of_traces C" by (metis (mono_tags, lifting) CollectI asm0(2) k_sem_def set_of_traces_def) qed have "(\i. (?l i, snd (?p i))) \ Q" proof (rule CHL_hyperpropE) show "CHL_hyperprop P Q (range ?p)" proof (rule CHL_hyperpropI) fix l p assume asm1: "(\i. p i \ range (\i. (snd (states i), snd (states' i)))) \ (\i. (l i, fst (p i))) \ P" then show "(\i. (l i, snd (p i))) \ Q" using CHL_hyperprop_def[of P Q "set_of_traces C"] \hypersat C (CHL_hyperprop P Q)\ \range (\i. (snd (states i), snd (states' i))) \ set_of_traces C\ hypersat_def subset_iff by blast qed show "(\i. (fst (states i), fst (snd (states i), snd (states' i)))) \ P" by (simp add: asm0(1)) fix i show "(snd (states i), snd (states' i)) \ range (\i. (snd (states i), snd (states' i)))" by blast qed moreover have "states' = (\i. (?l i, snd (?p i)))" proof (rule ext) fix i show "states' i = (fst (states i), snd (snd (states i), snd (states' i)))" by (metis asm0(2) k_sem_def prod.exhaust_sel sndI) qed ultimately show "states' \ Q" by auto qed next assume asm0: "CHL P C Q" have "CHL_hyperprop P Q (set_of_traces C)" proof (rule CHL_hyperpropI) fix l p assume asm1: "(\i. p i \ set_of_traces C) \ (\i. (l i, fst (p i))) \ P" show "(\i. (l i, snd (p i))) \ Q" using asm0 proof (rule CHLE) show "(\i. (l i, fst (p i))) \ P" by (simp add: asm1) show "k_sem C (\i. (l i, fst (p i))) (\i. (l i, snd (p i)))" proof (rule k_semI) fix i show "fst (l i, fst (p i)) = fst (l i, snd (p i)) \ \C, snd (l i, fst (p i))\ \ snd (l i, snd (p i))" using asm1 in_set_of_traces by fastforce qed qed qed then show "hypersat C (CHL_hyperprop P Q)" by (simp add: hypersat_def) qed theorem k_hypersafety_HL_hyperprop: fixes P :: "('i \ ('lvar, 'lval, 'pvar, 'pval) state) set" assumes "finite (UNIV :: 'i set)" and "card (UNIV :: 'i set) = k" shows "k_hypersafety k (CHL_hyperprop P Q)" proof (rule k_hypersafetyI) fix S assume "\ CHL_hyperprop P Q S" then obtain l p where asm0: "\i. p i \ S" "(\i. (l i, fst (p i))) \ P" "(\i. (l i, snd (p i))) \ Q" using CHL_hyperprop_def by blast let ?S = "range p" have "max_k k ?S" by (metis assms(1) assms(2) card_image_le finite_imageI max_k_def) moreover have "\S''. ?S \ S'' \ \ CHL_hyperprop P Q S''" by (meson asm0(2) asm0(3) CHL_hyperprop_def range_subsetD) ultimately show "\S'\S. max_k k S' \ (\S''. S' \ S'' \ \ CHL_hyperprop P Q S'')" by (meson asm0(1) image_subsetI) qed subsection \Incorrectness Logic~\cite{IncorrectnessLogic} or Reverse Hoare Logic~\cite{ReverseHL} (IL))\ -text \Definition 11\ +paragraph \Definition 18\ definition IL where "IL P C Q \ Q \ sem C P" lemma equiv_def_incorrectness: "IL P C Q \ (\l \'. (l, \') \ Q \ (\\. (l, \) \ P \ \C, \\ \ \'))" by (simp add: in_sem IL_def subset_iff) definition IL_hyperprop where "IL_hyperprop P Q S \ (\l \'. (l, \') \ Q \ (\\. (l, \) \ P \ (\, \') \ S))" lemma IL_hyperpropI: assumes "\l \'. (l, \') \ Q \ (\\. (l, \) \ P \ (\, \') \ S)" shows "IL_hyperprop P Q S" by (simp add: assms IL_hyperprop_def) -text \Proposition 11\ +paragraph \Proposition 5\ lemma IL_expresses_hyperproperties: "IL P C Q \ IL_hyperprop P Q (set_of_traces C)" (is "?A \ ?B") proof assume ?A show ?B proof (rule IL_hyperpropI) fix l \' assume asm0: "(l, \') \ Q" then obtain \ where "(l, \) \ P" "single_sem C \ \'" using \IL P C Q\ equiv_def_incorrectness by blast then show "\\. (l, \) \ P \ (\, \') \ set_of_traces C" using set_of_traces_def by auto qed next assume ?B have "Q \ sem C P" proof (rule subsetPairI) fix l \' assume "(l, \') \ Q" then obtain \ where "(\, \') \ set_of_traces C" "(l, \) \ P" by (meson \IL_hyperprop P Q (set_of_traces C)\ IL_hyperprop_def) then show "(l, \') \ sem C P" using in_set_of_traces_then_in_sem by blast qed then show ?A by (simp add: IL_def) qed lemma IL_consequence: assumes "IL P C Q" and "(l, \') \ Q" shows "\\. (l, \) \ P \ single_sem C \ \'" using assms(1) assms(2) equiv_def_incorrectness by blast -text \Proposition 4\ +paragraph \Proposition 6\ theorem encoding_IL: "IL P C Q \ (hyper_hoare_triple (under_approx P) C (under_approx Q))" (is "?A \ ?B") proof (rule iffI) show "?B \ ?A" proof - assume ?B then have "under_approx Q (sem C P)" by (simp add: hyper_hoare_triple_def under_approx_def) then show ?A by (simp add: IL_def under_approx_def) qed assume ?A then show ?B by (simp add: hyper_hoare_triple_def sem_monotonic IL_def under_approx_def subset_trans) qed lemma entailment_order_reverse_hoare: assumes "P \ P'" shows "entails (under_approx P') (under_approx P)" by (simp add: assms dual_order.trans entails_def under_approx_def) definition incorrectify where "incorrectify p = under_approx { \ |\. p \}" lemma incorrectifyI: assumes "\\. p \ \ \ \ S" shows "incorrectify p S" by (metis assms incorrectify_def mem_Collect_eq subsetI under_approx_def) lemma incorrectifyE: assumes "incorrectify p S" and "p \" shows "\ \ S" by (metis CollectI assms(1) assms(2) in_mono incorrectify_def under_approx_def) lemma simple_while_incorrectness: assumes "\n. hyper_hoare_triple (incorrectify (p n)) C (incorrectify (p (Suc n)))" shows "hyper_hoare_triple (incorrectify (p 0)) (While C) (incorrectify (\\. \n. p n \))" proof (rule consequence_rule) show "entails (incorrectify (p 0)) (incorrectify (p 0))" by (simp add: entailsI) show "hyper_hoare_triple (incorrectify (p 0)) (While C) (natural_partition (\n. incorrectify (p n)))" by (meson assms while_rule) have "entails (incorrectify (\\. \n. p n \)) (natural_partition (\n. incorrectify (p n)))" proof (rule entailsI) fix S assume asm0: "incorrectify (\\. \n. p n \) S" then have "under_approx { \ |\ n. p n \} S" by (metis incorrectify_def) let ?F = "\n. S" show "natural_partition (\n. incorrectify (p n)) S" proof (rule natural_partitionI) show "\n. incorrectify (p n) (?F n)" by (metis asm0 incorrectifyE incorrectifyI) show "S = \ (range ?F)" by simp qed qed show "entails (natural_partition (\n. incorrectify (p n))) (incorrectify (\\. \n. p n \))" proof (rule entailsI) fix S assume asm0: "natural_partition (\n. incorrectify (p n)) S" then obtain F where "S = (\n. F n)" "\n. incorrectify (p n) (F n)" using natural_partitionE by blast show "incorrectify (\\. \n. p n \) S" proof (rule incorrectifyI) fix \ assume "\n. p n \" then obtain n where "p n \" by blast then have "\ \ F n" by (meson \\n. incorrectify (p n) (F n)\ incorrectifyE) then show "\ \ S" using \S = \ (range F)\ by blast qed qed qed definition sat_for_l where "sat_for_l l P \ (\\. (l, \) \ P)" theorem incorrectness_hyperliveness: assumes "\l. sat_for_l l Q \ sat_for_l l P" shows "hyperliveness (IL_hyperprop P Q)" proof (rule hyperlivenessI) fix S let ?S = "S \ {(\, \') |\ \' l. (l, \') \ Q \ (l, \) \ P }" have "IL_hyperprop P Q ?S" proof (rule IL_hyperpropI) fix l \' assume asm0: "(l, \') \ Q" then obtain \ where "(l, \) \ P" by (meson assms sat_for_l_def) then show "\\. (l, \) \ P \ (\, \') \ ?S" using asm0 by auto qed then show "\S'. S \ S' \ IL_hyperprop P Q S'" by auto qed -subsection \Relational Incorrectness Logic~\cite{InsecurityLogic} (RIL)\ +subsection \k-Incorrectness Logic~\cite{InsecurityLogic} (k-IL)\ -text \Definition 11\ +text \RIL is the old name of k-IL.\ + +paragraph \Definition 19\ definition RIL where "RIL P C Q \ (\states' \ Q. \states \ P. k_sem C states states')" lemma RILI: assumes "\states'. states' \ Q \ (\states \ P. k_sem C states states')" shows "RIL P C Q" by (simp add: assms RIL_def) lemma RILE: assumes "RIL P C Q" and "states' \ Q" shows "\states \ P. k_sem C states states'" using assms(1) assms(2) RIL_def by blast definition RIL_hyperprop where "RIL_hyperprop P Q S \ (\l states'. (\i. (l i, states' i)) \ Q \ (\states. (\i. (l i, states i)) \ P \ (\i. (states i, states' i) \ S)))" lemma RIL_hyperpropI: assumes "\l states'. (\i. (l i, states' i)) \ Q \ (\states. (\i. (l i, states i)) \ P \ (\i. (states i, states' i) \ S))" shows "RIL_hyperprop P Q S" by (simp add: assms RIL_hyperprop_def) lemma RIL_hyperpropE: assumes "RIL_hyperprop P Q S" and "(\i. (l i, states' i)) \ Q" shows "\states. (\i. (l i, states i)) \ P \ (\i. (states i, states' i) \ S)" using assms(1) assms(2) RIL_hyperprop_def by blast lemma useful: "states' = (\i. ((fst \ states') i, (snd \ states') i))" proof (rule ext) fix i show "states' i = ((fst \ states') i, (snd \ states') i)" by auto qed -text \Proposition 12\ +paragraph \Proposition 17\ theorem RIL_expresses_hyperproperties: "hypersat C (RIL_hyperprop P Q) \ RIL P C Q" (is "?A \ ?B") proof assume ?A show ?B proof (rule RILI) fix states' assume asm0: "states' \ Q" then obtain states where asm0: "(\i. ((fst \ states') i, states i)) \ P \ (\i. (states i, (snd \ states') i) \ set_of_traces C)" using RIL_hyperpropE[of P Q "set_of_traces C" "fst \ states'" "snd \ states'"] \?A\ using hypersat_def by auto moreover have "k_sem C (\i. ((fst \ states') i, states i)) states'" proof (rule k_semI) fix i have "\C, snd ((fst \ states') i, states i)\ \ snd (states' i)" using calculation set_of_traces_def by auto then show "fst ((fst \ states') i, states i) = fst (states' i) \ \C, snd ((fst \ states') i, states i)\ \ snd (states' i)" by simp qed ultimately show "\states\P. k_sem C states states'" by fast qed next assume ?B have "RIL_hyperprop P Q (set_of_traces C)" proof (rule RIL_hyperpropI) fix l states' assume asm0: "(\i. (l i, states' i)) \ Q" then obtain states where "states \ P" "k_sem C states (\i. (l i, states' i))" using \RIL P C Q\ RILE by blast moreover have "(\i. (l i, (snd \ states) i)) = states" proof (rule ext) fix i show "(l i, (snd \ states) i) = states i" by (metis calculation(2) comp_apply fst_conv k_sem_def surjective_pairing) qed moreover have "\i. ((snd \ states) i, states' i) \ set_of_traces C" by (metis (mono_tags, lifting) calculation(2) comp_apply in_set_of_traces k_sem_def snd_conv) ultimately show "\states. (\i. (l i, states i)) \ P \ (\i. (states i, states' i) \ set_of_traces C)" by force qed then show ?A using hypersat_def by blast qed definition k_sat_for_l where "k_sat_for_l l P \ (\\. (\i. (l i, \ i)) \ P)" theorem RIL_hyperprop_hyperlive: assumes "\l. k_sat_for_l l Q \ k_sat_for_l l P" shows "hyperliveness (RIL_hyperprop P Q)" proof (rule hyperlivenessI) fix S have "RIL_hyperprop P Q UNIV" by (meson assms RIL_hyperpropI iso_tuple_UNIV_I k_sat_for_l_def) then show "\S'. S \ S' \ RIL_hyperprop P Q S'" by blast qed definition strong_pre_insec where "strong_pre_insec from_nat x c P S \ (\states \ P. (\i. fst (states i) x = from_nat i) \ (\r. \i. ((fst (states i))(c := r), snd (states i)) \ S)) \ (\states. (\i. states i \ S) \ (\i. fst (states i) x = from_nat i) \ (\i j. fst (states i) c = fst (states j) c) \ states \ P)" lemma strong_pre_insecI: assumes "\states. states \ P \ (\i. fst (states i) x = from_nat i) \ (\r. \i. ((fst (states i))(c := r), snd (states i)) \ S)" and "\states. (\i. states i \ S) \ (\i. fst (states i) x = from_nat i) \ (\i j. fst (states i) c = fst (states j) c) \ states \ P" shows "strong_pre_insec from_nat x c P S" by (simp add: assms(1) assms(2) strong_pre_insec_def) lemma strong_pre_insecE: assumes "strong_pre_insec from_nat x c P S" and "\i. states i \ S" and "\i. fst (states i) x = from_nat i" and "\i j. fst (states i) c = fst (states j) c" shows "states \ P" by (meson assms(1) assms(2) assms(3) assms(4) strong_pre_insec_def) definition pre_insec where "pre_insec from_nat x c P S \ (\states \ P. (\i. fst (states i) x = from_nat i) \ (\r. \i. ((fst (states i))(c := r), snd (states i)) \ S))" lemma pre_insecI: assumes "\states. states \ P \ (\i. fst (states i) x = from_nat i) \ (\r. \i. ((fst (states i))(c := r), snd (states i)) \ S)" shows "pre_insec from_nat x c P S" by (simp add: assms(1) pre_insec_def) lemma strong_pre_implies_pre: assumes "strong_pre_insec from_nat x c P S" shows "pre_insec from_nat x c P S" by (meson assms pre_insecI strong_pre_insec_def) lemma pre_insecE: assumes "pre_insec from_nat x c P S" and "states \ P" and "\i. fst (states i) x = from_nat i" shows "\r. \i. ((fst (states i))(c := r), snd (states i)) \ S" by (meson assms(1) assms(2) assms(3) pre_insec_def) definition post_insec where "post_insec from_nat x c Q S \ (\states \ Q. (\i. fst (states i) x = from_nat i) \ (\r. (\i. ((fst (states i))(c := r), snd (states i)) \ S)))" lemma post_insecE: assumes "post_insec from_nat x c Q S" and "states \ Q" and "\i. fst (states i) x = from_nat i" shows "\r. (\i. ((fst (states i))(c := r), snd (states i)) \ S)" by (meson assms(1) assms(2) assms(3) post_insec_def) lemma post_insecI: assumes "\states. states \ Q \ (\i. fst (states i) x = from_nat i) \ (\r. (\i. ((fst (states i))(c := r), snd (states i)) \ S))" shows "post_insec from_nat x c Q S" by (simp add: assms post_insec_def) lemma same_pre_post: "pre_insec from_nat x c Q S \ post_insec from_nat x c Q S" by (simp add: post_insec_def pre_insec_def) theorem can_be_sat: fixes x :: "'lvar" assumes "\l l' \. (\i. (l i, \ i)) \ P \ (\i. (l' i, \ i)) \ P" (* P does not depend on logical variables *) and "injective (indexify :: (('a \ ('pvar \ 'pval)) \ 'lval))" (* |lval| \ |PStates^k| *) and "x \ c" and "injective from_nat" shows "sat (strong_pre_insec from_nat x c (P :: ('a \ ('lvar \ 'lval) \ ('pvar \ 'pval)) set))" proof - let ?S = "\states\P. { (((fst (states i))(x := from_nat i))(c := indexify (\i. snd (states i))), snd (states i)) |i. True}" have "strong_pre_insec from_nat x c P ?S" proof (rule strong_pre_insecI) fix states assume asm0: "states \ P" "\i. fst (states i) x = from_nat i" define r where "r = indexify (\i. snd (states i))" have "\i. ((fst (states i))(c := r), snd (states i)) \ { (((fst (states i))(x := from_nat i))(c := indexify (\i. snd (states i))), snd (states i)) |i. True}" proof - fix i show "((fst (states i))(c := r), snd (states i)) \ { (((fst (states i))(x := from_nat i))(c := indexify (\i. snd (states i))), snd (states i)) |i. True}" using asm0(2) r_def by fastforce qed then show "\r. \i. ((fst (states i))(c := r), snd (states i)) \ ?S" by (meson UN_I asm0(1)) next fix states assume asm0: "\i. states i \ ?S" "\i. fst (states i) x = from_nat i" "\i j. fst (states i) c = fst (states j) c" let ?P = "\i states'. states' \ P \ states i \ { (((fst (states' i))(x := from_nat i))(c := indexify (\i. snd (states' i))), snd (states' i)) |i. True}" let ?states = "\i. SOME states'. ?P i states'" have r: "\i. ?P i (?states i)" proof - fix i show "?P i (?states i)" proof (rule someI_ex[of "?P i"]) show "\states'. states' \ P \ states i \ { (((fst (states' i))(x := from_nat i))(c := indexify (\i. snd (states' i))), snd (states' i)) |i. True}" using asm0(1) by fastforce qed qed moreover have rr: "\i. fst (states i) c = indexify (\j. snd (?states i j)) \ snd (?states i i) = snd (states i)" proof - fix i obtain j where j_def: "states i = (((fst ((?states i) j))(x := from_nat j))(c := indexify (\k. snd ((?states i) k))), snd ((?states i) j))" using r[of i] by blast then have r1: "snd (?states i j) = snd (states i)" by (metis (no_types, lifting) snd_conv) then have "from_nat i = from_nat j" by (metis (no_types, lifting) j_def asm0(2) assms(3) fst_conv fun_upd_same fun_upd_twist) then have "i = j" by (meson assms(4) injective_def) show "fst (states i) c = indexify (\j. snd (?states i j)) \ snd (?states i i) = snd (states i)" proof show "fst (states i) c = indexify (\j. snd (?states i j))" by (metis (no_types, lifting) j_def fst_conv fun_upd_same) show "snd (?states i i) = snd (states i)" using \i = j\ r1 by blast qed qed moreover have r0: "\i j. (\n. snd (?states i n)) = (\n. snd (?states j n))" proof - fix i j have "indexify (\n. snd (?states i n)) = indexify (\n. snd (?states j n))" using asm0(3) rr by fastforce then show "(\n. snd (?states i n)) = (\n. snd (?states j n))" by (meson assms(2) injective_def) qed obtain k :: 'a where "True" by blast then have "?states k \ P" using UN_iff[of _ "\states. {((fst (states i))(x := from_nat i, c := indexify (\i. snd (states i))), snd (states i)) |i. True}" P] asm0(1) someI_ex[of "\y. y \ P \ states k \ {((fst (y i))(x := from_nat i, c := indexify (\i. snd (y i))), snd (y i)) |i. True}"] by fast moreover have "\i. snd (?states k i) = snd (states i)" proof - fix i have "snd (?states i i) = snd (states i)" using rr by blast moreover have "(\n. snd (?states i n)) i = (\n. snd (?states k n)) i" by (metis r0) ultimately show "snd (?states k i) = snd (states i)" by auto qed moreover have "(\i. ((\i. fst (?states k i)) i, (\i. snd (states i)) i)) \ P \ (\i. ((\i. fst (states i)) i, (\i. snd (states i)) i)) \ P" using assms(1) by fast moreover have "(\i. ((\i. fst (states i)) i, (\i. snd (states i)) i)) = states" proof (rule ext) fix i show "(fst (states i), snd (states i)) = states i" by simp qed moreover have "(\i. ((\i. fst (?states k i)) i, (\i. snd (states i)) i)) = ?states k" proof (rule ext) fix i show "(\i. ((\i. fst (?states k i)) i, (\i. snd (states i)) i)) i = ?states k i" by (metis (no_types, lifting) calculation(4) prod.exhaust_sel) qed ultimately show "states \ P" by argo qed then show "sat (strong_pre_insec from_nat x c P)" by (meson sat_def) qed theorem encode_insec: assumes "injective from_nat" and "sat (strong_pre_insec from_nat x c (P :: ('a \ ('lvar \ 'lval) \ ('pvar \ 'pval)) set))" and "not_free_var_of P x \ not_free_var_of P c" and "not_free_var_of Q x \ not_free_var_of Q c" and "c \ x" shows "RIL P C Q \ \ {pre_insec from_nat x c P} C {post_insec from_nat x c Q}" (is "?A \ ?B") proof assume ?A show ?B proof (rule hyper_hoare_tripleI) fix S assume asm0: "pre_insec from_nat x c P S" show "post_insec from_nat x c Q (sem C S)" proof (rule post_insecI) fix states' assume asm1: "states' \ Q" "\i. fst (states' i) x = from_nat i" then obtain states where "states \ P" "k_sem C states states'" using \RIL P C Q\ RILE by blast then obtain r where asm2: "\i. ((fst (states i))(c := r), snd (states i)) \ S" using pre_insecE[of from_nat x c P S states] by (metis asm0 asm1(2) k_sem_def) then show "\r. \i. ((fst (states' i))(c := r), snd (states' i)) \ sem C S" by (metis (mono_tags, opaque_lifting) \k_sem C states states'\ k_sem_def single_step_then_in_sem) qed qed next assume asm0: ?B show ?A proof (rule RILI) fix states' assume asm1: "states' \ Q" obtain S where asm2: "strong_pre_insec from_nat x c P S" by (meson assms(2) sat_def) then have asm3: "post_insec from_nat x c Q (sem C S)" by (meson asm0 hyper_hoare_tripleE strong_pre_implies_pre) let ?states' = "\i. ((fst (states' i))(x := from_nat i), snd (states' i))" have "?states' \ Q" by (metis (no_types, lifting) asm1 assms(4) diff_by_update fstI not_free_var_of_def snd_conv) then obtain r where r_def: "\i. ((fst (?states' i))(c := r), snd (?states' i)) \ sem C S" using asm3 post_insecE[of from_nat x c Q] by fastforce let ?states = "\i. SOME \. ((fst (?states' i))(c := r), \) \ S \ single_sem C \ (snd (?states' i))" have asm4: "\i. ((fst (?states' i))(c := r), (?states i)) \ S \ single_sem C (?states i) (snd (?states' i))" proof - fix i have "\\. ((fst (?states' i))(c := r), \) \ S \ single_sem C \ (snd (?states' i))" by (metis r_def fst_conv in_sem snd_conv) then show "((fst (?states' i))(c := r), (?states i)) \ S \ single_sem C (?states i) (snd (?states' i))" using someI_ex[of "\\. ((fst (?states' i))(c := r), \) \ S \ single_sem C \ (snd (?states' i))"] by blast qed moreover have r0: "(\i. ((fst (?states' i))(c := r), (?states i))) \ P" using asm2 proof (rule strong_pre_insecE) fix i show "(\i. ((fst (?states' i))(c := r), (?states i))) i \ S" using calculation by blast show "fst ((\i. ((fst (?states' i))(c := r), (?states i))) i) x = from_nat i" using assms(5) by auto fix j show "fst ((\i. ((fst (?states' i))(c := r), (?states i))) i) c = fst ((\i. ((fst (?states' i))(c := r), (?states i))) j) c" by fastforce qed have r1: "(\i. (((fst (?states' i))(c := r))(x := fst (states' i) x), (?states i))) \ P" proof (rule not_free_var_ofE[of P x]) show "(\i. ((fst (?states' i))(c := r), (?states i))) \ P" using r0 by fastforce show "not_free_var_of P x" by (simp add: assms(3)) fix i show "differ_only_by (fst ((fst ((fst (states' i))(x := from_nat i), snd (states' i)))(c := r), ?states i)) (fst ((fst ((fst (states' i))(x := from_nat i), snd (states' i)))(c := r, x := fst (states' i) x), ?states i)) x" by (metis (mono_tags, lifting) diff_by_comm diff_by_update fst_conv) qed (auto) have "(\i. ((((fst (?states' i))(c := r))(x := fst (states' i) x))(c := fst (states' i) c), (?states i))) \ P" proof (rule not_free_var_ofE) show "(\i. (((fst (?states' i))(c := r))(x := fst (states' i) x), (?states i))) \ P" using r1 by fastforce show "not_free_var_of P c" by (simp add: assms(3)) fix i show "differ_only_by (fst ((fst ((fst (states' i))(x := from_nat i), snd (states' i)))(c := r, x := fst (states' i) x), ?states i)) (fst ((fst ((fst (states' i))(x := from_nat i), snd (states' i)))(c := r, x := fst (states' i) x, c := fst (states' i) c), ?states i)) c" by (metis (mono_tags, lifting) diff_by_comm diff_by_update fst_conv) qed (auto) moreover have "(\i. ((((fst (?states' i))(c := r))(x := fst (states' i) x))(c := fst (states' i) c), (?states i))) = (\i. (fst (states' i), (?states i)))" proof (rule ext) fix i show "(\i. ((((fst (?states' i))(c := r))(x := fst (states' i) x))(c := fst (states' i) c), (?states i))) i = (\i. (fst (states' i), (?states i))) i" by force qed moreover have "k_sem C (\i. (fst (states' i), (?states i))) states'" proof (rule k_semI) fix i show "(fst ((\i. (fst (states' i), (?states i))) i) = fst (states' i) \ single_sem C (snd ((\i. (fst (states' i), (?states i))) i)) (snd (states' i)))" using asm4 by auto qed ultimately show "\states\P. k_sem C states states'" by auto qed qed -text \Proposition 5\ +paragraph \Proposition 8\ theorem encoding_RIL: fixes x :: "'lvar" assumes "\l l' \. (\i. (l i, \ i)) \ P \ (\i. (l' i, \ i)) \ P" (* P does not depend on logical variables *) and "injective (indexify :: (('a \ ('pvar \ 'pval)) \ 'lval))" (* |lval| \ |PStates^k| *) and "c \ x" and "injective from_nat" and "not_free_var_of (P :: ('a \ ('lvar \ 'lval) \ ('pvar \ 'pval)) set) x \ not_free_var_of P c" and "not_free_var_of Q x \ not_free_var_of Q c" shows "RIL P C Q \ \ {pre_insec from_nat x c P} C {post_insec from_nat x c Q}" (is "?A \ ?B") proof (rule encode_insec) show "sat (strong_pre_insec from_nat x c (P :: ('a \ ('lvar \ 'lval) \ ('pvar \ 'pval)) set))" proof (rule can_be_sat) show "injective (indexify :: (('a \ ('pvar \ 'pval)) \ 'lval))" by (simp add: assms(2)) show "x \ c" using assms(3) by auto qed (auto simp add: assms) qed (auto simp add: assms) subsection \Forward Underapproximation (FU)\ text \As employed by Outcome Logic~\cite{OutcomeLogic}\ -text \Definition 12\ +paragraph \Definition 20\ definition FU where "FU P C Q \ (\l. \\. (l, \) \ P \ (\\'. single_sem C \ \' \ (l, \') \ Q))" lemma FUI: assumes "\\ l. (l, \) \ P \ (\\'. single_sem C \ \' \ (l, \') \ Q)" shows "FU P C Q" by (simp add: assms FU_def) definition encode_FU where "encode_FU P S \ P \ S \ {}" -text \Proposition 6\ +paragraph \Proposition 9\ theorem encoding_FU: "FU P C Q \ \ {encode_FU P} C {encode_FU Q}" (is "?A \ ?B") proof show "?B \ ?A" proof - assume ?B show ?A proof (rule FUI) fix \ l assume asm: "(l, \) \ P" then have "encode_FU P {(l, \)}" by (simp add: encode_FU_def) then have "Q \ sem C {(l, \)} \ {}" using \\ {encode_FU P} C {encode_FU Q}\ hyper_hoare_tripleE encode_FU_def by blast then obtain \' where "\' \ Q" "\' \ sem C {(l, \)}" by blast then show "\\'. single_sem C \ \' \ (l, \') \ Q" by (metis fst_conv in_sem prod.collapse singletonD snd_conv) qed qed assume ?A show ?B proof (rule hyper_hoare_tripleI) fix S assume "encode_FU P S" then obtain l \ where "(l, \) \ P \ S" by (metis Expressivity.encode_FU_def ex_in_conv surj_pair) then obtain \' where "single_sem C \ \' \ (l, \') \ Q" by (meson IntD1 \FU P C Q\ FU_def) then show "encode_FU Q (sem C S)" using Expressivity.encode_FU_def \(l, \) \ P \ S\ sem_def by fastforce qed qed definition hyperprop_FU where "hyperprop_FU P Q S \ (\l \. (l, \) \ P \ (\\'. (l, \') \ Q \ (\, \') \ S))" lemma hyperprop_FUI: assumes "\l \. (l, \) \ P \ (\\'. (l, \') \ Q \ (\, \') \ S)" shows "hyperprop_FU P Q S" by (simp add: hyperprop_FU_def assms) lemma hyperprop_FUE: assumes "hyperprop_FU P Q S" and "(l, \) \ P" shows "\\'. (l, \') \ Q \ (\, \') \ S" using hyperprop_FU_def assms(1) assms(2) by fastforce theorem FU_expresses_hyperproperties: "hypersat C (hyperprop_FU P Q) \ FU P C Q" (is "?A \ ?B") proof assume ?A show ?B proof (rule FUI) fix \ l assume "(l, \) \ P" then obtain \' where asm0: "(l, \') \ Q \ (\, \') \ set_of_traces C" by (meson \hypersat C (hyperprop_FU P Q)\ hyperprop_FUE hypersat_def) then show "\\'. (\C, \\ \ \') \ (l, \') \ Q" using in_set_of_traces by blast qed next assume ?B have "hyperprop_FU P Q (set_of_traces C)" proof (rule hyperprop_FUI) fix l \ assume asm0: "(l, \) \ P" then show "\\'. (l, \') \ Q \ (\, \') \ set_of_traces C" by (metis (mono_tags, lifting) CollectI \FU P C Q\ FU_def set_of_traces_def) qed then show ?A by (simp add: hypersat_def) qed theorem hyperliveness_hyperprop_FU: assumes "\l. sat_for_l l P \ sat_for_l l Q" shows "hyperliveness (hyperprop_FU P Q)" proof (rule hyperlivenessI) fix S show "\S'. S \ S' \ hyperprop_FU P Q S'" by (meson UNIV_I hyperprop_FU_def assms sat_for_l_def subsetI) qed -text \No relationship between incorrectness and forward underapproximation\ +paragraph \No relationship between incorrectness and forward underapproximation\ lemma incorrectness_does_not_imply_FU: assumes "injective from_nat" assumes "P = {(l, \) |\ l. \ x = from_nat (0 :: nat) \ \ x = from_nat 1}" and "Q = {(l, \) |\ l. \ x = from_nat 1}" and "C = Assume (\\. \ x = from_nat 1)" shows "IL P C Q" and "\ FU P C Q" proof - have "Q \ sem C P" proof (rule subsetPairI) fix l \ assume "(l, \) \ Q" then have "\ x = from_nat 1" using assms(3) by blast then have "(l, \) \ P" by (simp add: assms(2)) then show "(l, \) \ sem C P" by (simp add: \\ x = from_nat 1\ assms(4) sem_assume) qed then show "IL P C Q" by (simp add: IL_def) show "\ FU P C Q" proof (rule ccontr) assume "\ \ FU P C Q" then have "FU P C Q" by blast obtain \ where "\ x = from_nat 0" by simp then obtain l where "(l, \) \ P" using assms(2) by blast then obtain \' where "single_sem C \ \'" "(l, \') \ Q" by (meson \FU P C Q\ FU_def) then have "\' x = from_nat 0" using \\ x = from_nat 0\ assms(4) by blast then have "from_nat 0 = from_nat 1" using \\C, \\ \ \'\ assms(4) by force then show False using assms(1) injective_def[of from_nat] by auto qed qed lemma FU_does_not_imply_incorrectness: assumes "P = {(l, \) |\ l. \ x = from_nat (0 :: nat) \ \ x = from_nat 1}" and "Q = {(l, \) |\ l. \ x = from_nat 1}" assumes "injective from_nat" shows "\ IL Q Skip P" and "FU Q Skip P" proof - show "FU Q Skip P" proof (rule FUI) fix \ l assume "(l, \) \ Q" then show "\\'. (\Skip, \\ \ \') \ (l, \') \ P" using SemSkip assms(1) assms(2) by fastforce qed obtain \ where "\ x = from_nat 0" by simp then obtain l where "(l, \) \ P" using assms(1) by blast moreover have "\ x \ from_nat 1" by (metis \\ x = from_nat 0\ assms(3) injective_def one_neq_zero) then have "(l, \) \ Q" using assms(2) by blast ultimately show "\ IL Q Skip P" using IL_consequence by blast qed -subsection \Relational Forward Underapproximate logic\ +subsection \k-Forward Underapproximate logic\ -text \Definition 13\ +text \RFU is the old name of k-FU.\ + +paragraph \Definition 21\ definition RFU where "RFU P C Q \ (\states \ P. \states' \ Q. k_sem C states states')" lemma RFUI: assumes "\states. states \ P \ (\states' \ Q. k_sem C states states')" shows "RFU P C Q" by (simp add: assms RFU_def) lemma RFUE: assumes "RFU P C Q" and "states \ P" shows "\states' \ Q. k_sem C states states'" using assms(1) assms(2) RFU_def by blast definition encode_RFU where "encode_RFU from_nat x P S \ (\states \ P. (\i. states i \ S \ fst (states i) x = from_nat i))" -text \Proposition 7\ +paragraph \Proposition 11\ theorem encode_RFU: assumes "not_free_var_of P x" and "not_free_var_of Q x" and "injective from_nat" shows "RFU P C Q \ \ {encode_RFU from_nat x P} C {encode_RFU from_nat x Q}" (is "?A \ ?B") proof assume ?A show ?B proof (rule hyper_hoare_tripleI) fix S assume "encode_RFU from_nat x P S" then obtain states where asm0: "states \ P" "\i. states i \ S \ fst (states i) x = from_nat i" by (meson encode_RFU_def) then obtain states' where "states' \ Q" "k_sem C states states'" using \RFU P C Q\ RFUE by blast then have "\i. states' i \ sem C S \ fst (states' i) x = from_nat i" by (metis (mono_tags, lifting) asm0(2) in_sem k_sem_def prod.collapse) then show "encode_RFU from_nat x Q (sem C S)" by (meson \states' \ Q\ encode_RFU_def) qed next assume ?B show ?A proof (rule RFUI) fix states assume asm0: "states \ P" let ?states = "\i. ((fst (states i))(x := from_nat i), snd (states i))" have "?states \ P" using assms(1) proof (rule not_free_var_ofE) show "states \ P" using asm0 by simp fix i show "differ_only_by (fst (states i)) (fst ((fst (states i))(x := from_nat i), snd (states i))) x" using diff_by_comm diff_by_update by fastforce qed (auto) then have "encode_RFU from_nat x P (range ?states)" using encode_RFU_def by fastforce then have "encode_RFU from_nat x Q (sem C (range ?states))" using \\ {encode_RFU from_nat x P} C {encode_RFU from_nat x Q}\ hyper_hoare_tripleE by blast then obtain states' where states'_def: "states' \ Q" "\i. states' i \ sem C (range ?states) \ fst (states' i) x = from_nat i" by (meson encode_RFU_def) let ?states' = "\i. ((fst (states' i))(x := fst (states i) x), snd (states' i))" have "?states' \ Q" using assms(2) proof (rule not_free_var_ofE) show "states' \ Q" using \states' \ Q\ by simp fix i show "differ_only_by (fst (states' i)) (fst ((fst (states' i))(x := fst (states i) x), snd (states' i))) x" using diff_by_comm diff_by_update by fastforce qed (auto) moreover obtain to_pvar where to_pvar_def: "\i. to_pvar (from_nat i) = i" using assms(3) injective_then_exists_inverse by blast then have inj: "\i j. from_nat i = from_nat j \ i = j" by metis moreover have "k_sem C states ?states'" proof (rule k_semI) fix i obtain \ where "(fst (states' i), \) \ range (\i. ((fst (states i))(x := from_nat i), snd (states i))) \ \C, \\ \ snd (states' i)" using states'_def(2) in_sem by blast moreover have "fst (states' i) x = from_nat i" by (simp add: states'_def) then have r: "((fst (states (inv ?states (fst (states' i), \)))) (x := from_nat (inv ?states (fst (states' i), \))), snd (states (inv ?states (fst (states' i), \)))) = (fst (states' i), \)" by (metis (mono_tags, lifting) calculation f_inv_into_f) then have "single_sem C (snd (states i)) (snd (states' i))" using \fst (states' i) x = from_nat i\ calculation inj by fastforce moreover have "fst (?states i) = fst (states' i)" by (metis (mono_tags, lifting)r \fst (states' i) x = from_nat i\ fst_conv fun_upd_same inj) ultimately show "fst (states i) = fst ((fst (states' i))(x := fst (states i) x), snd (states' i)) \ \C, snd (states i)\ \ snd ((fst (states' i))(x := fst (states i) x), snd (states' i))" by (metis (mono_tags, lifting) fst_conv fun_upd_triv fun_upd_upd snd_conv) qed ultimately show "\states'\Q. k_sem C states states'" by blast qed qed definition RFU_hyperprop where "RFU_hyperprop P Q S \ (\l states. (\i. (l i, states i)) \ P \ (\states'. (\i. (l i, states' i)) \ Q \ (\i. (states i, states' i) \ S)))" lemma RFU_hyperpropI: assumes "\l states. (\i. (l i, states i)) \ P \ (\states'. (\i. (l i, states' i)) \ Q \ (\i. (states i, states' i) \ S))" shows "RFU_hyperprop P Q S" by (simp add: assms RFU_hyperprop_def) lemma RFU_hyperpropE: assumes "RFU_hyperprop P Q S" and "(\i. (l i, states i)) \ P" shows "\states'. (\i. (l i, states' i)) \ Q \ (\i. (states i, states' i) \ S)" using assms(1) assms(2) RFU_hyperprop_def by blast -text \Proposition 13\ +paragraph \Proposition 10\ theorem RFU_captures_hyperproperties: "hypersat C (RFU_hyperprop P Q) \ RFU P C Q" (is "?A \ ?B") proof assume ?A show ?B proof (rule RFUI) fix states assume "states \ P" moreover have "(\i. ((fst \ states) i, (snd \ states) i)) = states" by simp ultimately obtain states' where asm0: "(\i. ((fst \ states) i, states' i)) \ Q" "\i. ((snd \ states) i, states' i) \ set_of_traces C" using RFU_hyperpropE[of P Q "set_of_traces C" "fst \ states" "snd \ states"] using \hypersat C (RFU_hyperprop P Q)\ hypersat_def by auto moreover have "k_sem C states (\i. ((fst \ states) i, states' i))" proof (rule k_semI) fix i have "\C, snd (states i)\ \ states' i" using calculation(2) in_set_of_traces by fastforce then show "fst (states i) = fst ((fst \ states) i, states' i) \ \C, snd (states i)\ \ snd ((fst \ states) i, states' i)" by simp qed ultimately show "\states'\Q. k_sem C states states'" by fast qed next assume ?B have "RFU_hyperprop P Q (set_of_traces C)" proof (rule RFU_hyperpropI) fix l states assume "(\i. (l i, states i)) \ P" then obtain states' where asm0: "states' \ Q" "k_sem C (\i. (l i, states i)) states'" using \RFU P C Q\ RFUE by blast then have "\i. fst (states' i) = l i" by (simp add: k_sem_def) moreover have "(\i. (l i, (snd \ states') i)) = states'" proof (rule ext) fix i show "(l i, (snd \ states') i) = states' i" by (metis calculation comp_apply surjective_pairing) qed moreover have "\i. (states i, (snd \ states') i) \ set_of_traces C" proof - fix i show "(states i, (snd \ states') i) \ set_of_traces C" using asm0(2) comp_apply[of snd states'] in_set_of_traces k_sem_def[of C "\i. (l i, states i)" states'] snd_conv by fastforce qed ultimately show "\states'. (\i. (l i, states' i)) \ Q \ (\i. (states i, states' i) \ set_of_traces C)" using asm0(1) by force qed then show ?A by (simp add: hypersat_def) qed theorem hyperliveness_encode_RFU: assumes "\l. k_sat_for_l l P \ k_sat_for_l l Q" shows "hyperliveness (RFU_hyperprop P Q)" proof (rule hyperlivenessI) fix S have "RFU_hyperprop P Q UNIV" proof (rule RFU_hyperpropI) fix l states assume asm0: "(\i. (l i, states i)) \ P" then obtain states' where "(\i. (l i, states' i)) \ Q" by (metis assms k_sat_for_l_def) then show "\states'. (\i. (l i, states' i)) \ Q \ (\i. (states i, states' i) \ UNIV)" by blast qed then show "\S'. S \ S' \ RFU_hyperprop P Q S'" by blast qed -subsection \Relational Universal Existential (RUE)~\cite{RHLE}\ +subsection \k-Universal Existential (RUE)~\cite{RHLE}\ -text \Definition 14\ +text \RUE is the old name of k-UE.\ + +paragraph \Definition 22\ definition RUE where "RUE P C Q \ (\(\1, \2) \ P. \\1'. k_sem C \1 \1' \ (\\2'. k_sem C \2 \2' \ (\1', \2') \ Q))" lemma RUE_I: assumes "\\1 \2 \1'. (\1, \2) \ P \ k_sem C \1 \1' \ (\\2'. k_sem C \2 \2' \ (\1', \2') \ Q)" shows "RUE P C Q" using assms RUE_def by fastforce lemma RUE_E: assumes "RUE P C Q" and "(\1, \2) \ P" and "k_sem C \1 \1'" shows "\\2'. k_sem C \2 \2' \ (\1', \2') \ Q" using RUE_def assms(1) assms(2) assms(3) by blast -text \Hyperproperty\ +paragraph \Hyperproperty\ definition hyperprop_RUE where "hyperprop_RUE P Q S \ (\l1 l2 \1 \2 \1'. (\i. (l1 i, \1 i), \k. (l2 k, \2 k)) \ P \ (\i. (\1 i, \1' i) \ S) \ (\\2'. (\k. (\2 k, \2' k) \ S) \ (\i. (l1 i, \1' i), \k. (l2 k, \2' k)) \ Q))" lemma hyperprop_RUE_I: assumes "\l1 l2 \1 \2 \1'. (\i. (l1 i, \1 i), \k. (l2 k, \2 k)) \ P \ (\i. (\1 i, \1' i) \ S) \ (\\2'. (\k. (\2 k, \2' k) \ S) \ (\i. (l1 i, \1' i), \k. (l2 k, \2' k)) \ Q)" shows "hyperprop_RUE P Q S" using assms hyperprop_RUE_def[of P Q S] by force lemma hyperprop_RUE_E: assumes "hyperprop_RUE P Q S" and "(\i. (l1 i, \1 i), \k. (l2 k, \2 k)) \ P" and "\i. (\1 i, \1' i) \ S" shows "\\2'. (\k. (\2 k, \2' k) \ S) \ (\i. (l1 i, \1' i), \k. (l2 k, \2' k)) \ Q" using assms(1) assms(2) assms(3) hyperprop_RUE_def by blast -text \Proposition 14\ +paragraph \Proposition 12\ theorem RUE_express_hyperproperties: "RUE P C Q \ hypersat C (hyperprop_RUE P Q)" (is "?A \ ?B") proof assume asm0: ?A have "hyperprop_RUE P Q (set_of_traces C)" proof (rule hyperprop_RUE_I) fix l1 l2 \1 \2 \1' assume asm1: "(\i. (l1 i, \1 i), \k. (l2 k, \2 k)) \ P" "\i. (\1 i, \1' i) \ set_of_traces C" let ?x1 = "\i. (l1 i, \1 i)" let ?x2 = "\k. (l2 k, \2 k)" let ?x1' = "\i. (l1 i, \1' i)" have "\\2'. k_sem C (\k. (l2 k, \2 k)) \2' \ (?x1', \2') \ Q" using asm0 asm1(1) proof (rule RUE_E) show "k_sem C (\i. (l1 i, \1 i)) (\i. (l1 i, \1' i))" proof (rule k_semI) fix i have "single_sem C (\1 i) (\1' i)" using asm1(2) by (simp add: set_of_traces_def) then show "fst (l1 i, \1 i) = fst (l1 i, \1' i) \ \C, snd (l1 i, \1 i)\ \ snd (l1 i, \1' i)" by simp qed qed then obtain \2' where asm2: "k_sem C (\k. (l2 k, \2 k)) \2'" "(?x1', \2') \ Q" by blast let ?\2' = "\k. snd (\2' k)" have "\k. (\2 k, ?\2' k) \ set_of_traces C" by (metis (mono_tags, lifting) asm2(1) in_set_of_traces k_sem_def snd_conv) moreover have "(\k. (l2 k, ?\2' k)) = \2'" proof (rule ext) fix k show "(l2 k, snd (\2' k)) = \2' k" by (metis (mono_tags, lifting) asm2(1) fst_eqD k_sem_def surjective_pairing) qed ultimately show "\\2'. (\k. (\2 k, \2' k) \ set_of_traces C) \ (\i. (l1 i, \1' i), \k. (l2 k, \2' k)) \ Q" using asm2(2) by fastforce qed then show ?B by (simp add: hypersat_def) next assume ?B then have asm0: "hyperprop_RUE P Q (set_of_traces C)" by (simp add: hypersat_def) show ?A proof (rule RUE_I) fix \1 \2 \1' assume asm1: "(\1, \2) \ P" "k_sem C \1 \1'" then have "\i. fst (\1 i) = fst (\1' i)" by (simp add: k_sem_def) have "\\2'. (\k. (snd (\2 k), \2' k) \ set_of_traces C) \ (\i. (fst (\1 i), snd (\1' i)), \k. (fst (\2 k), \2' k)) \ Q" using asm0 proof (rule hyperprop_RUE_E[of P Q "set_of_traces C" "\i. fst (\1 i)" "\i. snd (\1 i)" "\k. fst (\2 k)" "\k. snd (\2 k)" "\i. snd (\1' i)"]) show "(\i. (fst (\1 i), snd (\1 i)), \k. (fst (\2 k), snd (\2 k))) \ P" by (simp add: asm1(1)) fix i show "(snd (\1 i), snd (\1' i)) \ set_of_traces C" by (metis (mono_tags, lifting) CollectI asm1(2) k_sem_def set_of_traces_def) qed then obtain \2' where asm2: "\k. (snd (\2 k), \2' k) \ set_of_traces C" "(\i. (fst (\1 i), snd (\1' i)), \k. (fst (\2 k), \2' k)) \ Q" by blast moreover have "k_sem C \2 (\k. (fst (\2 k), \2' k))" proof (rule k_semI) fix i show "fst (\2 i) = fst (fst (\2 i), \2' i) \ \C, snd (\2 i)\ \ snd (fst (\2 i), \2' i)" using calculation(1) in_set_of_traces by auto qed ultimately show "\\2'. k_sem C \2 \2' \ (\1', \2') \ Q" using \\i. fst (\1 i) = fst (\1' i)\ by auto qed qed definition is_type where "is_type type fn x t S \ \ (\i. \ i \ S \ fst (\ i) t = type \ fst (\ i) x = fn i)" lemma is_typeI: assumes "\i. \ i \ S" and "\i. fst (\ i) t = type" and "\i. fst (\ i) x = fn i" shows "is_type type fn x t S \" by (simp add: assms(1) assms(2) assms(3) is_type_def) lemma is_type_E: assumes "is_type type fn x t S \" shows "\ i \ S \ fst (\ i) t = type \ fst (\ i) x = fn i" by (meson assms is_type_def) definition encode_RUE_1 where "encode_RUE_1 fn fn1 fn2 x t P S \ (\k. \\ \ S. fst \ x = fn2 k \ fst \ t = fn 2) \ (\\ \'. is_type (fn 1) fn1 x t S \ \ is_type (fn 2) fn2 x t S \' \ (\, \') \ P)" lemma encode_RUE_1_I: assumes "\k. \\ \ S. fst \ x = fn2 k \ fst \ t = fn 2" and "\\ \'. is_type (fn 1) fn1 x t S \ \ is_type (fn 2) fn2 x t S \' \ (\, \') \ P" shows "encode_RUE_1 fn fn1 fn2 x t P S" by (simp add: assms(1) assms(2) encode_RUE_1_def) lemma encode_RUE_1_E1: assumes "encode_RUE_1 fn fn1 fn2 x t P S" shows "\\ \ S. fst \ x = fn2 k \ fst \ t = fn 2" by (meson assms encode_RUE_1_def) lemma encode_RUE_1_E2: assumes "encode_RUE_1 fn fn1 fn2 x t P S" and "is_type (fn 1) fn1 x t S \" and "is_type (fn 2) fn2 x t S \'" shows "(\, \') \ P" by (meson assms encode_RUE_1_def) definition encode_RUE_2 where "encode_RUE_2 fn fn1 fn2 x t Q S \ (\\. is_type (fn 1) fn1 x t S \ \ (\\'. is_type (fn 2) fn2 x t S \' \ (\, \') \ Q))" lemma encode_RUE_2I: assumes "\\. is_type (fn 1) fn1 x t S \ \ (\\'. is_type (fn 2) fn2 x t S \' \ (\, \') \ Q)" shows "encode_RUE_2 fn fn1 fn2 x t Q S" by (simp add: assms encode_RUE_2_def) lemma encode_RUE_2_E: assumes "encode_RUE_2 fn fn1 fn2 x t Q S" and "is_type (fn 1) fn1 x t S \" shows "\\'. is_type (fn 2) fn2 x t S \' \ (\, \') \ Q" by (meson assms(1) assms(2) encode_RUE_2_def) definition differ_only_by_set where "differ_only_by_set vars a b \ (\x. x \ vars \ a x = b x)" definition differ_only_by_lset where "differ_only_by_lset vars a b \ (\i. snd (a i) = snd (b i) \ differ_only_by_set vars (fst (a i)) (fst (b i)))" lemma differ_only_by_lsetI: assumes "\i. snd (a i) = snd (b i)" and "\i. differ_only_by_set vars (fst (a i)) (fst (b i))" shows "differ_only_by_lset vars a b" using assms(1) assms(2) differ_only_by_lset_def by blast definition not_in_free_vars_double where "not_in_free_vars_double vars P \ (\\ \'. differ_only_by_lset vars (fst \) (fst \') \ differ_only_by_lset vars (snd \) (snd \') \ (\ \ P \ \' \ P))" lemma not_in_free_vars_doubleE: assumes "not_in_free_vars_double vars P" and "differ_only_by_lset vars (fst \) (fst \')" and "differ_only_by_lset vars (snd \) (snd \')" and "\ \ P" shows "\' \ P" by (meson assms not_in_free_vars_double_def) -text \Proposition 8\ +paragraph \Proposition 13\ theorem encoding_RUE: assumes "injective fn \ injective fn1 \ injective fn2" and "t \ x" and "injective (fn :: nat \ 'a)" and "injective fn1" and "injective fn2" and "not_in_free_vars_double {x, t} P" and "not_in_free_vars_double {x, t} Q" shows "RUE P C Q \ \ {encode_RUE_1 fn fn1 fn2 x t P} C {encode_RUE_2 fn fn1 fn2 x t Q}" (is "?A \ ?B") proof assume asm0: ?A show ?B proof (rule hyper_hoare_tripleI) fix S assume asm1: "encode_RUE_1 fn fn1 fn2 x t P S" show "encode_RUE_2 fn fn1 fn2 x t Q (sem C S)" proof (rule encode_RUE_2I) fix \1' assume asm2: "is_type (fn 1) fn1 x t (sem C S) \1'" let ?\2 = "\k. SOME \'. \'\S \ fst \' x = fn2 k \ fst \' t = fn 2" have r2: "\k. ?\2 k \S \ fst (?\2 k) x = fn2 k \ fst (?\2 k) t = fn 2" proof - fix k show "?\2 k \S \ fst (?\2 k) x = fn2 k \ fst (?\2 k) t = fn 2" proof (rule someI_ex) show "\xa. xa \ S \ fst xa x = fn2 k \ fst xa t = fn 2" by (meson asm1 encode_RUE_1_E1) qed qed let ?\1 = "\i. SOME \. (fst (\1' i), \) \ S \ single_sem C \ (snd (\1' i))" have r1: "\i. (fst (\1' i), ?\1 i) \ S \ single_sem C (?\1 i) (snd (\1' i))" proof - fix i show "(fst (\1' i), ?\1 i) \ S \ single_sem C (?\1 i) (snd (\1' i))" proof (rule someI_ex[of "\\. (fst (\1' i), \) \ S \ single_sem C \ (snd (\1' i))"]) show "\\. (fst (\1' i), \) \ S \ \C, \\ \ snd (\1' i)" by (meson asm2 in_sem is_type_def) qed qed have "(\i. (fst (\1' i), ?\1 i), ?\2) \ P" using asm1 proof (rule encode_RUE_1_E2) show "is_type (fn 1) fn1 x t S (\i. (fst (\1' i), ?\1 i))" using asm2 fst_conv is_type_def[of "fn 1" fn1 x t S] is_type_def[of "fn 1" fn1 x t "sem C S"] r1 by force show "is_type (fn 2) fn2 x t S ?\2" by (simp add: is_type_def r2) qed moreover have "\\2'. k_sem C ?\2 \2' \ (\1', \2') \ Q" using asm0 proof (rule RUE_E) show "(\i. (fst (\1' i), ?\1 i), ?\2) \ P" using calculation by auto show "k_sem C (\i. (fst (\1' i), SOME \. (fst (\1' i), \) \ S \ \C, \\ \ snd (\1' i))) \1'" by (simp add: k_sem_def r1) qed then obtain \2' where \2'_def: "k_sem C ?\2 \2' \ (\1', \2') \ Q" by blast then have "is_type (fn 2) fn2 x t (sem C S) \2'" using in_sem[of _ C S] k_semE[of C ?\2 \2'] prod.collapse r2 is_type_def[of "fn 2" fn2 x t S] is_type_def[of "fn 2" fn2 x t "sem C S"] by (metis (no_types, lifting)) then show "\\2'. is_type (fn 2) fn2 x t (sem C S) \2' \ (\1', \2') \ Q" using \2'_def by blast qed qed next assume asm0: "\ {encode_RUE_1 fn fn1 fn2 x t P} C {encode_RUE_2 fn fn1 fn2 x t Q}" show ?A proof (rule RUE_I) fix \1 \2 \1' assume asm1: "(\1, \2) \ P" "k_sem C \1 \1'" let ?\1 = "\i. ((((fst (\1 i))(t := fn 1))(x := fn1 i)), snd (\1 i))" let ?\2 = "\k. ((((fst (\2 k))(t := fn 2))(x := fn2 k)), snd (\2 k))" let ?S1 = "{ ?\1 i |i. True }" let ?S2 = "{ ?\2 k |k. True }" let ?S = "?S1 \ ?S2" let ?\1' = "\i. ((((fst (\1' i))(t := fn 1))(x := fn1 i)), snd (\1' i))" have "encode_RUE_2 fn fn1 fn2 x t Q (sem C ?S)" using asm0 proof (rule hyper_hoare_tripleE) show "encode_RUE_1 fn fn1 fn2 x t P ?S" proof (rule encode_RUE_1_I) fix k let ?\ = "((((fst (\2 k))(t := fn 2))(x := fn2 k)), snd (\2 k))" have "?\ \ ?S2" by auto moreover have "fst ?\ x = fn2 k" by simp moreover have "fst ?\ t = fn 2" by (simp add: assms(2)) ultimately show "\\\?S1 \ ?S2. fst \ x = fn2 k \ fst \ t = fn 2" by blast next fix \ \' assume asm2: "is_type (fn (1 :: nat)) fn1 x t (?S1 \ ?S2) \ \ is_type (fn 2) fn2 x t (?S1 \ ?S2) \'" moreover have r1: "\i. \ i = ((fst (\1 i))(t := fn 1, x := fn1 i), snd (\1 i))" proof - fix i have "fst (\ i) t = fn 1" by (meson calculation is_type_def) moreover have "\ i \ ?S1" proof (rule ccontr) assume "\ \ i \ ?S1" moreover have "\ i \ ?S1 \ ?S2" using asm2 is_type_def[of "fn 1" fn1 x t] by (metis (no_types, lifting)) ultimately have "\ i \ ?S2" by simp then have "fst (\ i) t = fn 2" using assms(2) by auto then show False by (metis Suc_1 Suc_eq_numeral \fst (\ i) t = fn 1\ assms(3) injective_def numeral_One one_neq_zero pred_numeral_simps(1)) qed then obtain j where "\ i = ((fst (\1 j))(t := fn 1, x := fn1 j), snd (\1 j))" by blast moreover have "i = j" by (metis (mono_tags, lifting) asm2 assms(4) calculation(2) fst_conv fun_upd_same injective_def is_type_def) ultimately show "\ i = ((fst (\1 i))(t := fn 1, x := fn1 i), snd (\1 i))" by blast qed moreover have "\i. \' i = ((fst (\2 i))(t := fn 2, x := fn2 i), snd (\2 i))" proof - fix i have "fst (\' i) t = fn 2" by (meson calculation is_type_def) moreover have "\' i \ ?S2" proof (rule ccontr) assume "\ \' i \ ?S2" moreover have "\' i \ ?S1 \ ?S2" using asm2 is_type_def[of "fn 2" fn2 x t] by (metis (no_types, lifting)) ultimately have "\' i \ ?S1" by simp then have "fst (\' i) t = fn 1" using assms(2) by auto then show False by (metis Suc_1 Suc_eq_numeral \fst (\' i) t = fn 2\ assms(3) injective_def numeral_One one_neq_zero pred_numeral_simps(1)) qed then obtain j where "\' i = ((fst (\2 j))(t := fn 2, x := fn2 j), snd (\2 j))" by blast moreover have "i = j" by (metis (mono_tags, lifting) asm2 assms(5) calculation(2) fst_conv fun_upd_same injective_def is_type_def) ultimately show "\' i = ((fst (\2 i))(t := fn 2, x := fn2 i), snd (\2 i))" by blast qed moreover have "(?\1, ?\2) \ P" using assms(6) proof (rule not_in_free_vars_doubleE) show "(\1, \2) \ P" by (simp add: asm1(1)) show "differ_only_by_lset {x, t} (fst (\1, \2)) (fst (?\1, ?\2))" by (rule differ_only_by_lsetI) (simp_all add: differ_only_by_set_def) show "differ_only_by_lset {x, t} (snd (\1, \2)) (snd (?\1, ?\2))" by (rule differ_only_by_lsetI) (simp_all add: differ_only_by_set_def) qed ultimately show "(\, \') \ P" by presburger qed qed then have "\\'. is_type (fn 2) fn2 x t (sem C ?S) \' \ (?\1', \') \ Q" proof (rule encode_RUE_2_E) show "is_type (fn 1) fn1 x t (sem C ?S) ?\1'" proof (rule is_typeI) fix i show "fst ((fst (\1' i))(t := fn 1, x := fn1 i), snd (\1' i)) t = fn 1" by (simp add: assms(2)) show " ((fst (\1' i))(t := fn 1, x := fn1 i), snd (\1' i)) \ sem C ?S" using UnI1[of _ ?S1 ?S2] asm1(2) k_semE[of C \1 \1' i] single_step_then_in_sem[of C "snd (\1 i)" "snd (\1' i)" _ ?S] by force qed (auto) qed then obtain \2' where r: "is_type (fn 2) fn2 x t (sem C ?S) \2' \ (?\1', \2') \ Q" by blast let ?\2' = "\k. ((fst (\2' k))(x := fst (\2 k) x, t := fst (\2 k) t), snd (\2' k))" have "(\1', ?\2') \ Q" using assms(7) proof (rule not_in_free_vars_doubleE) show "(?\1', \2') \ Q" using r by blast show "differ_only_by_lset {x, t} (fst (?\1', \2')) (fst (\1', ?\2'))" by (rule differ_only_by_lsetI) (simp_all add: differ_only_by_set_def) show "differ_only_by_lset {x, t} (snd (?\1', \2')) (snd (\1', ?\2'))" by (rule differ_only_by_lsetI) (simp_all add: differ_only_by_set_def) qed moreover have "k_sem C \2 ?\2'" proof (rule k_semI) fix i obtain y where y_def: "y \ ?S" "fst y = fst (\2' i)" "single_sem C (snd y) (snd (\2' i))" using r in_sem[of "\2' i" C ?S] is_type_E[of "fn 2" fn2 x t "sem C ?S" \2' i] by (metis (no_types, lifting) fst_conv snd_conv) then have "fst y t = fn 2" by (metis (no_types, lifting) is_type_def r) moreover have "fn 1 \ fn 2" by (metis Suc_1 assms(3) injective_def n_not_Suc_n) then have "y \ ?S1" using assms(2) calculation by fastforce then have "y \ ?S2" using y_def(1) by blast show "fst (\2 i) = fst ((fst (\2' i))(x := fst (\2 i) x, t := fst (\2 i) t), snd (\2' i)) \ \C, snd (\2 i)\ \ snd ((fst (\2' i))(x := fst (\2 i) x, t := fst (\2 i) t), snd (\2' i))" proof have r1: "\2' i \ sem C ?S \ fst (\2' i) t = fn 2 \ fst (\2' i) x = fn2 i" proof (rule is_type_E[of "fn 2" fn2 x t "sem C ?S" \2' i]) show "is_type (fn 2) fn2 x t (sem C ?S) \2'" using r by blast qed then obtain \ where "(fst (\2' i), \) \ ?S" "single_sem C \ (snd (\2' i))" by (meson in_sem) then have "(fst (\2' i), \) \ ?S2" using r1 \fn 1 \ fn 2\ assms(2) by fastforce then obtain k where "fst (\2' i) = (fst (\2 k))(t := fn 2, x := fn2 k)" and "\ = snd (\2 k)" by blast then have "k = i" by (metis r1 assms(5) fun_upd_same injective_def) then show "\C, snd (\2 i)\ \ snd ((fst (\2' i))(x := fst (\2 i) x, t := fst (\2 i) t), snd (\2' i))" using \\C, \\ \ snd (\2' i)\ \\ = snd (\2 k)\ by auto show "fst (\2 i) = fst ((fst (\2' i))(x := fst (\2 i) x, t := fst (\2 i) t), snd (\2' i))" by (simp add: \fst (\2' i) = (fst (\2 k))(t := fn 2, x := fn2 k)\ \k = i\) qed qed ultimately show "\\2'. k_sem C \2 \2' \ (\1', \2') \ Q" by blast qed qed subsection \Program Refinement\ lemma sem_assign_single: "sem (Assign x e) {(l, \)} = {(l, \(x := e \))}" (is "?A = ?B") proof show "?A \ ?B" proof (rule subsetPairI) fix la \' assume "(la, \') \ sem (Assign x e) {(l, \)}" then show "(la, \') \ {(l, \(x := e \))}" by (metis (mono_tags, lifting) in_sem prod.sel(1) prod.sel(2) single_sem_Assign_elim singleton_iff) qed show "?B \ ?A" by (simp add: SemAssign in_sem) qed definition refinement where "refinement C1 C2 \ (set_of_traces C1 \ set_of_traces C2)" definition not_free_var_stmt where "not_free_var_stmt x C \ (\\ \' v. (\, \') \ set_of_traces C \ (\(x := v), \'(x := v)) \ set_of_traces C) \ (\\ \'. single_sem C \ \' \ \ x = \' x)" lemma not_free_var_stmtE_1: assumes "not_free_var_stmt x C" and "(\, \') \ set_of_traces C" shows "(\(x := v), \'(x := v)) \ set_of_traces C" using assms(1) assms(2) not_free_var_stmt_def by force lemma not_free_in_sem_same_val: assumes "not_free_var_stmt x C" and "single_sem C \ \'" shows "\ x = \' x" using assms(1) assms(2) not_free_var_stmt_def by fastforce lemma not_free_in_sem_equiv: assumes "not_free_var_stmt x C" and "single_sem C \ \'" shows "single_sem C (\(x := v)) (\'(x := v))" by (meson assms(1) assms(2) in_set_of_traces not_free_var_stmtE_1) -text \Example 4\ +paragraph \Example 3\ + +lemma rewrite_if_commute: + assumes "\ { P } If C1 C2 {Q}" + shows "\ { P } If C2 C1 {Q}" + by (metis assms hyper_hoare_triple_def sem_if sup_commute) + theorem encoding_refinement: fixes P :: "(('lvar \ 'lval) \ ('pvar \ 'pval)) set \ bool" assumes "(a :: 'pval) \ b" (* and x free var *) and "P = (\S. card S = 1)" and "Q = (\S. \\\S. snd \ x = a \ (fst \, (snd \)(x := b)) \ S)" and "not_free_var_stmt x C1" and "not_free_var_stmt x C2" shows "refinement C1 C2 \ \ { P } If (Seq (Assign (x :: 'pvar) (\_. a)) C1) (Seq (Assign x (\_. b)) C2) { Q }" (is "?A \ ?B") proof assume ?A show ?B proof (rule hyper_hoare_tripleI) fix S assume "P (S :: (('lvar \ 'lval) \ ('pvar \ 'pval)) set)" then obtain \ l where asm0: "S = {(l, \)}" by (metis assms(2) card_1_singletonE surj_pair) let ?C = "If (Seq (Assign x (\_. a)) C1) (Seq (Assign x (\_. b)) C2)" let ?a = "(l, \(x := a))" let ?b = "(l, \(x := b))" have if_sem: "sem ?C S = sem C1 {?a} \ sem C2 {?b}" by (simp add: asm0 sem_assign_single sem_if sem_seq) then have "\\. \ \ sem ?C S \ snd \ x = a \ (fst \, (snd \)(x := b)) \ sem ?C S" proof - fix \ assume asm1: "\ \ sem ?C S" "snd \ x = a" have "\ \ sem C1 {?a}" proof (rule ccontr) assume "\ \ sem C1 {(l, \(x := a))}" then have "\ \ sem C2 {(l, \(x := b))}" using if_sem asm1(1) by force then have "snd \ x = b" using assms(5) fun_upd_same in_sem not_free_in_sem_same_val[of x C2 "\(x := b)" "snd \"] singletonD snd_conv by metis then show False using asm1(2) assms(1) by blast qed then have "(\(x := a), snd \) \ set_of_traces C1" by (simp add: in_sem set_of_traces_def) then have "(\(x := a), snd \) \ set_of_traces C2" using \refinement C1 C2\ refinement_def by fastforce then have "((\(x := a))(x := b), (snd \)(x := b)) \ set_of_traces C2" by (meson assms(5) not_free_var_stmtE_1) then have "single_sem C2 (\(x := b)) ((snd \)(x := b))" by (simp add: set_of_traces_def) then have "(fst \, (snd \)(x := b)) \ sem C2 {?b}" by (metis \\ \ sem C1 {(l, \(x := a))}\ fst_eqD in_sem singleton_iff snd_eqD) then show "(fst \, (snd \)(x := b)) \ sem ?C S" by (simp add: if_sem) qed then show "Q (sem ?C S)" using assms(3) by blast qed next assume asm0: ?B have "set_of_traces C1 \ set_of_traces C2" proof (rule subsetPairI) fix \ \' assume asm1: "(\, \') \ set_of_traces C1" obtain l S where "(S :: (('lvar \ 'lval) \ ('pvar \ 'pval)) set) = { (l, \) }" by simp let ?a = "(l, \(x := a))" let ?b = "(l, \(x := b))" let ?C = "If (Seq (Assign (x :: 'pvar) (\_. a)) C1) (Seq (Assign x (\_. b)) C2)" have "Q (sem ?C S)" proof (rule hyper_hoare_tripleE) show "P S" by (simp add: \S = {(l, \)}\ assms(2)) show ?B using asm0 by simp qed moreover have "(l, \'(x := a)) \ sem ?C S" proof - have "single_sem (Seq (Assign x (\_. a)) C1) \ (\'(x := a))" by (meson SemAssign SemSeq asm1 assms(4) in_set_of_traces not_free_in_sem_equiv) then show ?thesis by (simp add: SemIf1 \S = {(l, \)}\ in_sem) qed then have "(l, \'(x := b)) \ sem ?C S" using assms(3) calculation by fastforce moreover have "(l, \'(x := b)) \ sem (Seq (Assign x (\_. b)) C2) S" proof (rule ccontr) assume "\ (l, \'(x := b)) \ sem (Seq (Assign x (\_. b)) C2) S" then have "(l, \'(x := b)) \ sem (Seq (Assign x (\_. a)) C1) S" using calculation(2) sem_if by auto then have "(l, \'(x := b)) \ sem C1 {?a}" by (simp add: \S = {(l, \)}\ sem_assign_single sem_seq) then show False using assms(1) assms(4) fun_upd_same in_sem not_free_in_sem_same_val[of x C1 "\(x := a)" "\'(x := b)"] singletonD snd_conv by metis qed then have "single_sem (Seq (Assign x (\_. b)) C2) \ (\'(x := b))" by (simp add: \S = {(l, \)}\ in_sem) then have "single_sem C2 (\(x := b)) (\'(x := b))" by blast then have "(\(x := b), \'(x := b)) \ set_of_traces C2" by (simp add: set_of_traces_def) then have "((\(x := b))(x := \ x), (\'(x := b))(x := \ x)) \ set_of_traces C2" by (meson assms(5) not_free_var_stmtE_1) then show "(\, \') \ set_of_traces C2" by (metis asm1 assms(4) fun_upd_triv fun_upd_upd in_set_of_traces not_free_in_sem_same_val) qed then show ?A by (simp add: refinement_def) qed +paragraph \Necessary Preconditions\ + +definition NC where + "NC P C Q \ (\\ \' l. (l, \') \ Q \ (\C, \\ \ \') \ (l, \) \ P)" + +lemma NC_I: + assumes "\\ \' l. (l, \') \ Q \ (\C, \\ \ \') \ (l, \) \ P" + shows "NC P C Q" + by (simp add: NC_def assms) + +definition backwards_sem where + "backwards_sem C S' = { (l, \) |l \ \'. (l, \') \ S' \ \C, \\ \ \'}" + +lemma equiv_def_NC: + "NC P C Q \ backwards_sem C Q \ P" (is "?A \ ?B") +proof + assume ?A + then show ?B + using CollectD NC_def backwards_sem_def[of C Q] subsetI[of "backwards_sem C Q" P] by fastforce +next + assume ?B + then show ?A + using NC_def backwards_sem_def by fastforce +qed + +lemma equiv_def_FU: + "FU P C Q \ P \ backwards_sem C Q" (is "?A \ ?B") +proof + show "?A \ ?B" + using FU_def[of P C Q] backwards_sem_def[of C Q] + by fastforce + show "?B \ ?A" + using FU_def[of P C Q] backwards_sem_def[of C Q] by auto +qed + +lemma encoding_NC_in_HL_1: + "NC P C Q \ \ { (\S. S = -P) } C { (\S. S \ Q = {} ) }" (is "?A \ ?B") +proof + assume ?A + show ?B + proof (rule hyper_hoare_tripleI) + fix S assume asm0: "S = - P" + then show "sem C S \ Q = {}" + by (metis (no_types, lifting) ComplD NC_def \NC P C Q\ disjoint_iff in_sem prod.collapse) + qed +next + assume ?B + show ?A + proof (rule NC_I) + fix \ \' l + assume asm0: "(l, \') \ Q \ \C, \\ \ \'" + then have "(l, \') \ sem C (-P)" + using \\ {(\S. S = - P)} C {\S. S \ Q = {}}\ hyper_hoare_tripleE by fastforce + then have "(l, \) \ -P" + by (meson asm0 single_step_then_in_sem) + then show "(l, \) \ P" + by simp + qed +qed + + end \ No newline at end of file diff --git a/thys/HyperHoareLogic/Language.thy b/thys/HyperHoareLogic/Language.thy --- a/thys/HyperHoareLogic/Language.thy +++ b/thys/HyperHoareLogic/Language.thy @@ -1,341 +1,356 @@ section \Language and Semantics\ -text \In this file, we formalize the programming language from section III, and the extended states -and semantics from section IV of the paper~\cite{HyperHoareLogic}. We also prove the useful properties described by Lemma 1.\ +text \In this file, we formalize concepts from section 3: +- Program states and programming language (definition 1) +- Big-step semantics (figure 2) +- Extended states (definition 2) +- Extended semantics (definition 4) and some useful properties (lemma 1)\ theory Language imports Main begin subsection Language text \Definition 1\ type_synonym ('var, 'val) pstate = "'var \ 'val" -text \Definition 2\ - type_synonym ('var, 'val) bexp = "('var, 'val) pstate \ bool" type_synonym ('var, 'val) exp = "('var, 'val) pstate \ 'val" datatype ('var, 'val) stmt = Assign 'var "('var, 'val) exp" - | Seq "('var, 'val) stmt" "('var, 'val) stmt" - | If "('var, 'val) stmt" "('var, 'val) stmt" + | Seq "('var, 'val) stmt" "('var, 'val) stmt" (infixl ";;" 60) + | If "('var, 'val) stmt" "('var, 'val) stmt" \\Non-deterministic choice\ | Skip - | Havoc 'var + | Havoc 'var \\Non-deterministic assignment\ | Assume "('var, 'val) bexp" - | While "('var, 'val) stmt" + | While "('var, 'val) stmt" \\Non-deterministic loop\ subsection Semantics text \Figure 2\ inductive single_sem :: "('var, 'val) stmt \ ('var, 'val) pstate \ ('var, 'val) pstate \ bool" ("\_, _\ \ _" [51,0] 81) where SemSkip: "\Skip, \\ \ \" | SemAssign: "\Assign var e, \\ \ \(var := (e \))" | SemSeq: "\ \C1, \\ \ \1; \C2, \1\ \ \2 \ \ \Seq C1 C2, \\ \ \2" | SemIf1: "\C1, \\ \ \1 \ \If C1 C2, \\ \ \1" | SemIf2: "\C2, \\ \ \2 \ \If C1 C2, \\ \ \2" | SemHavoc: "\Havoc var, \\ \ \(var := v)" | SemAssume: "b \ \ \Assume b, \\ \ \" | SemWhileIter: "\ \C, \\ \ \' ; \While C, \'\ \ \'' \ \ \While C, \\ \ \''" | SemWhileExit: "\While C, \\ \ \" inductive_cases single_sem_Seq_elim[elim!]: "\Seq C1 C2, \\ \ \'" inductive_cases single_sem_Skip_elim[elim!]: "\Skip, \\ \ \'" inductive_cases single_sem_While_elim: "\While C, \\ \ \'" inductive_cases single_sem_If_elim[elim!]: "\If C1 C2, \\ \ \'" inductive_cases single_sem_Assume_elim[elim!]: "\Assume b, \\ \ \'" inductive_cases single_sem_Assign_elim[elim!]: "\Assign x e, \\ \ \'" inductive_cases single_sem_Havoc_elim[elim!]: "\Havoc x, \\ \ \'" -section \Extended States and Extended Semantics\ +subsection \Extended States and Extended Semantics\ -text \Definition 3\ +text \Definition 2: Extended states\ type_synonym ('lvar, 'lval, 'pvar, 'pval) state = "('lvar \ 'lval) \ ('pvar, 'pval) pstate" -text \Definition 5\ +text \Definition 4: Extended semantics\ definition sem :: "('pvar, 'pval) stmt \ ('lvar, 'lval, 'pvar, 'pval) state set \ ('lvar, 'lval, 'pvar, 'pval) state set" where "sem C S = { (l, \') |\' \ l. (l, \) \ S \ \C, \\ \ \' }" lemma in_sem: "\ \ sem C S \ (\\. (fst \, \) \ S \ single_sem C \ (snd \))" (is "?A \ ?B") proof assume ?A then obtain \' \ l where "\ = (l, \')" "(l, \) \ S \ \C, \\ \ \'" using sem_def[of C S] by auto then show ?B by auto next show "?B \ ?A" by (metis (mono_tags, lifting) CollectI prod.collapse sem_def) qed -text \Useful properties\ +text \Lemma 1: Useful properties of the extended semantics\ lemma sem_seq: "sem (Seq C1 C2) S = sem C2 (sem C1 S)" (is "?A = ?B") proof show "?A \ ?B" proof fix x2 assume "x2 \ ?A" then obtain x0 where "(fst x2, x0) \ S" "single_sem (Seq C1 C2) x0 (snd x2)" by (metis in_sem) then obtain x1 where "single_sem C1 x0 x1" "single_sem C2 x1 (snd x2)" using single_sem_Seq_elim[of C1 C2 x0 "snd x2"] by blast then show "x2 \ ?B" by (metis \(fst x2, x0) \ S\ fst_conv in_sem snd_conv) qed show "?B \ ?A" proof fix x2 assume "x2 \ ?B" then obtain x1 where "(fst x2, x1) \ sem C1 S" "single_sem C2 x1 (snd x2)" by (metis in_sem) then obtain x0 where "(fst x2, x0) \ S" "single_sem C1 x0 x1" by (metis fst_conv in_sem snd_conv) then have "single_sem (Seq C1 C2) x0 (snd x2)" by (simp add: SemSeq \\C2, x1\ \ snd x2\) then show "x2 \ ?A" by (meson \(fst x2, x0) \ S\ in_sem) qed qed lemma sem_skip: "sem Skip S = S" using single_sem_Skip_elim SemSkip in_sem[of _ Skip S] by fastforce lemma sem_union: "sem C (S1 \ S2) = sem C S1 \ sem C S2" (is "?A = ?B") proof show "?A \ ?B" proof fix x assume "x \ ?A" then obtain y where "(fst x, y) \ S1 \ S2" "single_sem C y (snd x)" using in_sem by blast then show "x \ ?B" by (metis Un_iff in_sem) qed show "?B \ ?A" proof fix x assume "x \ ?B" show "x \ ?A" proof (cases "x \ sem C S1") case True then show ?thesis by (metis IntD2 Un_Int_eq(3) in_sem) next case False then show ?thesis by (metis Un_iff \x \ sem C S1 \ sem C S2\ in_sem) qed qed qed lemma sem_union_general: "sem C (\x. f x) = (\x. sem C (f x))" (is "?A = ?B") proof show "?A \ ?B" proof fix b assume "b \ ?A" then obtain a where "a \ (\x. f x)" "fst a = fst b" "single_sem C (snd a) (snd b)" by (metis fst_conv in_sem snd_conv) then obtain x where "a \ f x" by blast then have "b \ sem C (f x)" by (metis \\C, snd a\ \ snd b\ \fst a = fst b\ in_sem surjective_pairing) then show "b \ ?B" by blast qed show "?B \ ?A" proof fix y assume "y \ ?B" then obtain x where "y \ sem C (f x)" by blast then show "y \ ?A" by (meson UN_I in_sem iso_tuple_UNIV_I) qed qed lemma sem_monotonic: assumes "S \ S'" shows "sem C S \ sem C S'" by (metis assms sem_union subset_Un_eq) lemma subsetPairI: assumes "\l \. (l, \) \ A \ (l, \) \ B" shows "A \ B" by (simp add: assms subrelI) lemma sem_if: "sem (If C1 C2) S = sem C1 S \ sem C2 S" (is "?A = ?B") proof show "?A \ ?B" proof (rule subsetPairI) fix l y assume "(l, y) \ ?A" then obtain x where "(l, x) \ S" "single_sem (If C1 C2) x y" by (metis fst_conv in_sem snd_conv) then show "(l, y) \ ?B" using single_sem_If_elim UnI1 UnI2 in_sem by (metis fst_conv snd_conv) qed show "?B \ ?A" using SemIf1 SemIf2 in_sem by (metis (no_types, lifting) Un_subset_iff subsetI) qed lemma sem_assume: "sem (Assume b) S = { (l, \) |l \. (l, \) \ S \ b \ }" (is "?A = ?B") proof show "?A \ ?B" proof (rule subsetPairI) fix l y assume "(l, y) \ ?A" then obtain x where "(l, x) \ S" "single_sem (Assume b) x y" using in_sem by (metis fst_conv snd_conv) then show "(l, y) \ ?B" using single_sem_Assume_elim by blast qed show "?B \ ?A" proof (rule subsetPairI) fix l \ assume asm0: "(l, \) \ {(l, \) |l \. (l, \) \ S \ b \}" then have "(l, \) \ S" "b \" by simp_all then show "(l, \) \ sem (Assume b) S" by (metis SemAssume fst_eqD in_sem snd_eqD) qed qed lemma while_then_reaches: assumes "(single_sem C)\<^sup>*\<^sup>* \ \''" shows "single_sem (While C) \ \''" using assms proof (induct rule: converse_rtranclp_induct) case base then show ?case by (simp add: SemWhileExit) next case (step y z) then show ?case using SemWhileIter by blast qed lemma in_closure_then_while: assumes "single_sem C' \ \''" shows "\C. C' = While C \ (single_sem C)\<^sup>*\<^sup>* \ \''" using assms proof (induct rule: single_sem.induct) case (SemWhileIter \ C' \' \'') then show ?case by (metis (no_types, lifting) rtranclp.rtrancl_into_rtrancl rtranclp.rtrancl_refl rtranclp_trans stmt.inject(6)) next case (SemWhileExit \ C') then show ?case by blast qed (auto) theorem loop_equiv: "single_sem (While C) \ \' \ (single_sem C)\<^sup>*\<^sup>* \ \'" using in_closure_then_while while_then_reaches by blast fun iterate_sem where "iterate_sem 0 _ S = S" | "iterate_sem (Suc n) C S = sem C (iterate_sem n C S)" lemma in_iterate_then_in_trans: assumes "(l, \'') \ iterate_sem n C S" shows "\\. (l, \) \ S \ (single_sem C)\<^sup>*\<^sup>* \ \''" using assms proof (induct n arbitrary: \'' S) case 0 then show ?case using iterate_sem.simps(1) by blast next case (Suc n) then show ?case using in_sem rtranclp.rtrancl_into_rtrancl by (metis (mono_tags, lifting) fst_conv iterate_sem.simps(2) snd_conv) qed lemma reciprocal: assumes "(single_sem C)\<^sup>*\<^sup>* \ \''" and "(l, \) \ S" shows "\n. (l, \'') \ iterate_sem n C S" using assms proof (induct rule: rtranclp_induct) case base then show ?case using iterate_sem.simps(1) by blast next case (step y z) then obtain n where "(l, y) \ iterate_sem n C S" by blast then show ?case using in_sem iterate_sem.simps(2) step.hyps(2) by (metis fst_eqD snd_eqD) qed lemma union_iterate_sem_trans: "(l, \'') \ (\n. iterate_sem n C S) \ (\\. (l, \) \ S \ (single_sem C)\<^sup>*\<^sup>* \ \'')" (is "?A \ ?B") using in_iterate_then_in_trans reciprocal by auto lemma sem_while: "sem (While C) S = (\n. iterate_sem n C S)" (is "?A = ?B") proof show "?A \ ?B" proof (rule subsetPairI) fix l y assume "(l, y) \ ?A" then obtain x where x_def: "(l, x) \ S" "(single_sem C)\<^sup>*\<^sup>* x y" using in_closure_then_while in_sem by (metis fst_eqD snd_conv) then have "single_sem (While C) x y" using while_then_reaches by blast then show "(l, y) \ ?B" by (metis x_def union_iterate_sem_trans) qed show "?B \ ?A" proof (rule subsetPairI) fix l y assume "(l, y) \ ?B" then obtain x where "(l, x) \ S" "(single_sem C)\<^sup>*\<^sup>* x y" using union_iterate_sem_trans by blast then show "(l, y) \ ?A" using in_sem while_then_reaches by fastforce qed qed lemma assume_sem: "sem (Assume b) S = Set.filter (b \ snd) S" (is "?A = ?B") proof show "?A \ ?B" proof (rule subsetPairI) fix l \ assume asm0: "(l, \) \ sem (Assume b) S" then show "(l, \) \ Set.filter (b \ snd) S" by (metis comp_apply fst_conv in_sem member_filter single_sem_Assume_elim snd_conv) qed show "?B \ ?A" by (metis (mono_tags, opaque_lifting) SemAssume comp_apply in_sem member_filter prod.collapse subsetI) qed lemma sem_split_general: "sem C (\x. F x) = (\x. sem C (F x))" (is "?A = ?B") proof show "?A \ ?B" proof (rule subsetPairI) fix l \' assume asm0: "(l, \') \ sem C (\ (range F))" then obtain x \ where "(l, \) \ F x" "single_sem C \ \'" by (metis (no_types, lifting) UN_iff fst_conv in_sem snd_conv) then show "(l, \') \ (\x. sem C (F x))" using asm0 sem_union_general by blast qed show "?B \ ?A" by (simp add: SUP_least Sup_upper sem_monotonic) qed +fun written_vars where + "written_vars (Assign x _) = {x}" +| "written_vars (Havoc x) = {x}" +| "written_vars (C1;; C2) = written_vars C1 \ written_vars C2" +| "written_vars (If C1 C2) = written_vars C1 \ written_vars C2" +| "written_vars (While C) = written_vars C" +| "written_vars _ = {}" + +lemma written_vars_not_modified: + assumes "single_sem C \ \'" + and "x \ written_vars C" + shows "\ x = \' x" + using assms + by (induct rule: single_sem.induct) auto end \ No newline at end of file diff --git a/thys/HyperHoareLogic/Logic.thy b/thys/HyperHoareLogic/Logic.thy --- a/thys/HyperHoareLogic/Logic.thy +++ b/thys/HyperHoareLogic/Logic.thy @@ -1,1018 +1,857 @@ section \Hyper Hoare Logic\ -text \In this file, we define concepts from the logic (section IV): hyper-assertions, hyper-triples, -and the syntactic rules. We also prove soundness (theorem 1), completeness (theorem 2), the ability -to disprove hyper-triples in the logic (theorem 4), and the synchronized if rule from appendix C.\ +text \This file contains technical results from sections 3 and 5: +- Hyper-assertions (definition 3) +- Hyper-triples (definition 5) +- Core rules of Hyper Hoare Logic (figure 2) +- Soundness of the core rules (theorem 1) +- Completeness of the core rules (theorem 2) +- Ability to disprove hyper-triples (theorem 5)\ theory Logic imports Language begin -text \Definition 4\ +text \Definition 3\ type_synonym 'a hyperassertion = "('a set \ bool)" definition entails where "entails A B \ (\S. A S \ B S)" +lemma entails_refl: + "entails A A" + by (simp add: entails_def) + lemma entailsI: assumes "\S. A S \ B S" shows "entails A B" by (simp add: assms entails_def) lemma entailsE: assumes "entails A B" and "A x" shows "B x" by (meson assms(1) assms(2) entails_def) lemma bientails_equal: assumes "entails A B" and "entails B A" shows "A = B" proof (rule ext) fix S show "A S = B S" by (meson assms(1) assms(2) entailsE) qed lemma entails_trans: assumes "entails A B" and "entails B C" shows "entails A C" by (metis assms(1) assms(2) entails_def) definition setify_prop where "setify_prop b = { (l, \) |l \. b \}" lemma sem_assume_setify: "sem (Assume b) S = S \ setify_prop b" (is "?A = ?B") proof - have "\l \. (l, \) \ ?A \ (l, \) \ ?B" proof - fix l \ have "(l, \) \ ?A \ (l, \) \ S \ b \" by (simp add: assume_sem) then show "(l, \) \ ?A \ (l, \) \ ?B" by (simp add: setify_prop_def) qed then show ?thesis by auto qed definition over_approx :: "'a set \ 'a hyperassertion" where "over_approx P S \ S \ P" definition lower_closed :: "'a hyperassertion \ bool" where "lower_closed P \ (\S S'. P S \ S' \ S \ P S')" lemma over_approx_lower_closed: "lower_closed (over_approx P)" by (metis (full_types) lower_closed_def order_trans over_approx_def) definition under_approx :: "'a set \ 'a hyperassertion" where "under_approx P S \ P \ S" definition upper_closed :: "'a hyperassertion \ bool" where "upper_closed P \ (\S S'. P S \ S \ S' \ P S')" lemma under_approx_upper_closed: "upper_closed (under_approx P)" by (metis (no_types, lifting) order.trans under_approx_def upper_closed_def) definition closed_by_union :: "'a hyperassertion \ bool" where "closed_by_union P \ (\S S'. P S \ P S' \ P (S \ S'))" lemma closed_by_unionI: assumes "\a b. P a \ P b \ P (a \ b)" shows "closed_by_union P" by (simp add: assms closed_by_union_def) lemma closed_by_union_over: "closed_by_union (over_approx P)" by (simp add: closed_by_union_def over_approx_def) lemma closed_by_union_under: "closed_by_union (under_approx P)" by (simp add: closed_by_union_def sup.coboundedI1 under_approx_def) definition conj where "conj P Q S \ P S \ Q S" +lemma entail_conj: + assumes "entails A B" + shows "entails A (conj A B)" + by (metis (full_types) assms conj_def entails_def) + +lemma entail_conj_weaken: + "entails (conj A B) A" + by (simp add: conj_def entails_def) + definition disj where "disj P Q S \ P S \ Q S" definition exists :: "('c \ 'a hyperassertion) \ 'a hyperassertion" where "exists P S \ (\x. P x S)" definition forall :: "('c \ 'a hyperassertion) \ 'a hyperassertion" where "forall P S \ (\x. P x S)" lemma over_inter: "entails (over_approx (P \ Q)) (conj (over_approx P) (over_approx Q))" by (simp add: conj_def entails_def over_approx_def) lemma over_union: "entails (disj (over_approx P) (over_approx Q)) (over_approx (P \ Q))" by (metis disj_def entailsI le_supI1 le_supI2 over_approx_def) lemma under_union: "entails (under_approx (P \ Q)) (disj (under_approx P) (under_approx Q))" by (simp add: disj_def entails_def under_approx_def) lemma under_inter: "entails (conj (under_approx P) (under_approx Q)) (under_approx (P \ Q))" by (simp add: conj_def entails_def le_infI1 under_approx_def) - -text \Notation 1\ +text \Definition 6: Operator \\\\ definition join :: "'a hyperassertion \ 'a hyperassertion \ 'a hyperassertion" where "join A B S \ (\SA SB. A SA \ B SB \ S = SA \ SB)" definition general_join :: "('b \ 'a hyperassertion) \ 'a hyperassertion" where "general_join f S \ (\F. S = (\x. F x) \ (\x. f x (F x)))" +lemma general_joinI: + assumes "S = (\x. F x)" + and "\x. f x (F x)" + shows "general_join f S" + using assms(1) assms(2) general_join_def by blast lemma join_closed_by_union: assumes "closed_by_union Q" shows "join Q Q = Q" proof fix S show "join Q Q S \ Q S" by (metis assms closed_by_union_def join_def sup_idem) qed lemma entails_join_entails: assumes "entails A1 B1" and "entails A2 B2" shows "entails (join A1 A2) (join B1 B2)" proof (rule entailsI) fix S assume "join A1 A2 S" then obtain S1 S2 where "A1 S1" "A2 S2" "S = S1 \ S2" by (metis join_def) then show "join B1 B2 S" by (metis assms(1) assms(2) entailsE join_def) qed -text \Notation 2\ +text \Definition 7: Operator \\\ (for \x \ X\)\ definition natural_partition where "natural_partition I S \ (\F. S = (\n. F n) \ (\n. I n (F n)))" lemma natural_partitionI: assumes "S = (\n. F n)" and "\n. I n (F n)" shows "natural_partition I S" using assms(1) assms(2) natural_partition_def by blast lemma natural_partitionE: assumes "natural_partition I S" obtains F where "S = (\n. F n)" "\n. I n (F n)" by (meson assms natural_partition_def) subsection \Rules of the Logic\ -text \Rules from figure 3\ +text \Core rules from figure 2\ inductive syntactic_HHT :: "(('lvar, 'lval, 'pvar, 'pval) state hyperassertion) \ ('pvar, 'pval) stmt \ (('lvar, 'lval, 'pvar, 'pval) state hyperassertion) \ bool" ("\ {_} _ {_}" [51,0,0] 81) where RuleSkip: "\ {P} Skip {P}" | RuleCons: "\ entails P P' ; entails Q' Q ; \ {P'} C {Q'} \ \ \ {P} C {Q}" | RuleSeq: "\ \ {P} C1 {R} ; \ {R} C2 {Q} \ \ \ {P} (Seq C1 C2) {Q}" | RuleIf: "\ \ {P} C1 {Q1} ; \ {P} C2 {Q2} \ \ \ {P} (If C1 C2) {join Q1 Q2}" | RuleWhile: "\ \n. \ {I n} C {I (Suc n)} \ \ \ {I 0} (While C) {natural_partition I}" | RuleAssume: "\ { (\S. P (Set.filter (b \ snd) S)) } (Assume b) {P}" | RuleAssign: "\ { (\S. P { (l, \(x := e \)) |l \. (l, \) \ S }) } (Assign x e) {P}" | RuleHavoc: "\ { (\S. P { (l, \(x := v)) |l \ v. (l, \) \ S }) } (Havoc x) {P}" | RuleExistsSet: "\\x::('lvar, 'lval, 'pvar, 'pval) state set. \ {P x} C {Q x}\ \ \ {exists P} C {exists Q}" subsection \Soundness\ -text \Definition 6: Hyper-Triples\ +text \Definition 5: Hyper-Triples\ definition hyper_hoare_triple ("\ {_} _ {_}" [51,0,0] 81) where "\ {P} C {Q} \ (\S. P S \ Q (sem C S))" lemma hyper_hoare_tripleI: assumes "\S. P S \ Q (sem C S)" shows "\ {P} C {Q}" by (simp add: assms hyper_hoare_triple_def) lemma hyper_hoare_tripleE: assumes "\ {P} C {Q}" and "P S" shows "Q (sem C S)" using assms(1) assms(2) hyper_hoare_triple_def by metis lemma consequence_rule: assumes "entails P P'" and "entails Q' Q" and "\ {P'} C {Q'}" shows "\ {P} C {Q}" by (metis (no_types, opaque_lifting) assms(1) assms(2) assms(3) entails_def hyper_hoare_triple_def) lemma skip_rule: "\ {P} Skip {P}" by (simp add: hyper_hoare_triple_def sem_skip) lemma assume_rule: "\ { (\S. P (Set.filter (b \ snd) S)) } (Assume b) {P}" proof (rule hyper_hoare_tripleI) fix S assume "P (Set.filter (b \ snd) S)" then show "P (sem (Assume b) S)" by (simp add: assume_sem) qed lemma seq_rule: assumes "\ {P} C1 {R}" and "\ {R} C2 {Q}" shows "\ {P} Seq C1 C2 {Q}" using assms(1) assms(2) hyper_hoare_triple_def sem_seq by metis lemma if_rule: assumes "\ {P} C1 {Q1}" and "\ {P} C2 {Q2}" shows "\ {P} If C1 C2 {join Q1 Q2}" by (metis (full_types) assms(1) assms(2) hyper_hoare_triple_def join_def sem_if) lemma sem_assign: "sem (Assign x e) S = {(l, \(x := e \)) |l \. (l, \) \ S}" (is "?A = ?B") proof show "?A \ ?B" proof (rule subsetPairI) fix l \' assume "(l, \') \ sem (Assign x e) S" then obtain \ where "(l, \) \ S" "single_sem (Assign x e) \ \'" by (metis fst_eqD in_sem snd_conv) then show "(l, \') \ {(l, \(x := e \)) |l \. (l, \) \ S}" by blast qed show "?B \ ?A" proof (rule subsetPairI) fix l \' assume "(l, \') \ ?B" then obtain \ where "\' = \(x := e \)" "(l, \) \ S" by blast then show "(l, \') \ ?A" by (metis SemAssign fst_eqD in_sem snd_conv) qed qed lemma assign_rule: "\ { (\S. P { (l, \(x := e \)) |l \. (l, \) \ S }) } (Assign x e) {P}" proof (rule hyper_hoare_tripleI) fix S assume "P {(l, \(x := e \)) |l \. (l, \) \ S}" then show "P (sem (Assign x e) S)" using sem_assign by metis qed lemma sem_havoc: "sem (Havoc x) S = {(l, \(x := v)) |l \ v. (l, \) \ S}" (is "?A = ?B") proof show "?A \ ?B" proof (rule subsetPairI) fix l \' assume "(l, \') \ sem (Havoc x) S" then obtain \ where "(l, \) \ S" "single_sem (Havoc x) \ \'" by (metis fst_eqD in_sem snd_conv) then show "(l, \') \ {(l, \(x := v)) |l \ v. (l, \) \ S}" by blast qed show "?B \ ?A" proof (rule subsetPairI) fix l \' assume "(l, \') \ ?B" then obtain \ v where "\' = \(x := v)" "(l, \) \ S" by blast then show "(l, \') \ ?A" by (metis SemHavoc fst_eqD in_sem snd_conv) qed qed lemma havoc_rule: "\ { (\S. P { (l, \(x := v)) |l \ v. (l, \) \ S }) } (Havoc x) {P}" proof (rule hyper_hoare_tripleI) fix S assume "P { (l, \(x := v)) |l \ v. (l, \) \ S }" then show "P (sem (Havoc x) S)" using sem_havoc by metis qed text \Loops\ lemma indexed_invariant_then_power: assumes "\n. hyper_hoare_triple (I n) C (I (Suc n))" and "I 0 S" shows "I n (iterate_sem n C S)" using assms proof (induct n arbitrary: S) next case (Suc n) then have "I n (iterate_sem n C S)" by blast then have "I (Suc n) (sem C (iterate_sem n C S))" using Suc.prems(1) hyper_hoare_tripleE by blast then show ?case by (simp add: Suc.hyps Suc.prems(1)) qed (auto) +lemma indexed_invariant_then_power_bounded: + assumes "\m. m < n \ hyper_hoare_triple (I m) C (I (Suc m))" + and "I 0 S" + shows "I n (iterate_sem n C S)" + using assms +proof (induct n arbitrary: S) +next + case (Suc n) + then have "I n (iterate_sem n C S)" + using less_Suc_eq by presburger + then have "I (Suc n) (sem C (iterate_sem n C S))" + using Suc.prems(1) hyper_hoare_tripleE by blast + then show ?case + by (simp add: Suc.hyps Suc.prems(1)) +qed (auto) + lemma while_rule: assumes "\n. hyper_hoare_triple (I n) C (I (Suc n))" shows "hyper_hoare_triple (I 0) (While C) (natural_partition I)" proof (rule hyper_hoare_tripleI) fix S assume asm0: "I 0 S" show "natural_partition I (sem (While C) S)" proof (rule natural_partitionI) show "sem (While C) S = \ (range (\n. iterate_sem n C S))" by (simp add: sem_while) fix n show "I n (iterate_sem n C S)" by (simp add: asm0 assms indexed_invariant_then_power) qed qed - -text \Additional rules\ - -lemma empty_pre: - "hyper_hoare_triple (\_. False) C QQ" - by (simp add: hyper_hoare_triple_def) - -lemma full_post: - "hyper_hoare_triple P C (\_. True)" - by (simp add: hyper_hoare_triple_def) - - -lemma rule_join: - assumes "\ {P} C {Q}" - and "hyper_hoare_triple P' C Q'" - shows "hyper_hoare_triple (join P P') C (join Q Q')" -proof (rule hyper_hoare_tripleI) - fix S assume asm0: "join P P' S" - then obtain S1 S2 where "S = S1 \ S2" "P S1" "P' S2" - by (metis join_def) - then have "sem C S = sem C S1 \ sem C S2" - using sem_union by auto - then show "join Q Q' (sem C S)" - by (metis \P S1\ \P' S2\ assms(1) assms(2) hyper_hoare_tripleE join_def) -qed - -lemma rule_general_join: - assumes "\x. \ {P x} C {Q x}" - shows "hyper_hoare_triple (general_join P) C (general_join Q)" -proof (rule hyper_hoare_tripleI) - fix S assume "general_join P S" - then obtain F where asm0: "S = (\x. F x)" "\x. P x (F x)" - by (meson general_join_def) - have "sem C S = (\x. sem C (F x))" - by (simp add: asm0(1) sem_split_general) - moreover have "\x. Q x (sem C (F x))" - using asm0(2) assms hyper_hoare_tripleE by blast - ultimately show "general_join Q (sem C S)" - by (metis general_join_def) -qed - -lemma rule_conj: - assumes "\ {P} C {Q}" - and "hyper_hoare_triple P' C Q'" - shows "hyper_hoare_triple (conj P P') C (conj Q Q')" -proof (rule hyper_hoare_tripleI) - fix S assume "Logic.conj P P' S" - then show "Logic.conj Q Q' (sem C S)" - by (metis assms(1) assms(2) conj_def hyper_hoare_tripleE) -qed - -text \Generalization\ - -lemma rule_forall: - assumes "\x. \ {P x} C {Q x}" - shows "hyper_hoare_triple (forall P) C (forall Q)" - by (metis assms forall_def hyper_hoare_triple_def) - -lemma rule_disj: - assumes "\ {P} C {Q}" - and "\ {P'} C {Q'}" - shows "hyper_hoare_triple (disj P P') C (disj Q Q')" - by (metis assms(1) assms(2) disj_def hyper_hoare_triple_def) - -text \Generalization\ - lemma rule_exists: assumes "\x. \ {P x} C {Q x}" shows "\ {exists P} C {exists Q}" by (metis assms exists_def hyper_hoare_triple_def) -corollary variant_if_rule: - assumes "hyper_hoare_triple P C1 Q" - and "hyper_hoare_triple P C2 Q" - and "closed_by_union Q" - shows "hyper_hoare_triple P (If C1 C2) Q" - by (metis assms(1) assms(2) assms(3) if_rule join_closed_by_union) - - -text \Simplifying the rule\ - -definition stable_by_infinite_union :: "'a hyperassertion \ bool" where - "stable_by_infinite_union I \ (\F. (\S \ F. I S) \ I (\S \ F. S))" - -lemma stable_by_infinite_unionE: - assumes "stable_by_infinite_union I" - and "\S. S \ F \ I S" - shows "I (\S \ F. S)" - using assms(1) assms(2) stable_by_infinite_union_def by blast - -lemma stable_by_union_and_constant_then_I: - assumes "\n. I n = I'" - and "stable_by_infinite_union I'" - shows "natural_partition I = I'" -proof (rule ext) - fix S show "natural_partition I S = I' S" - proof - show "I' S \ natural_partition I S" - proof - - assume "I' S" - show "natural_partition I S" - proof (rule natural_partitionI) - show "S = \ (range (\n. S))" - by simp - fix n show "I n S" - by (simp add: \I' S\ assms(1)) - qed - qed - assume asm0: "natural_partition I S" - then obtain F where "S = (\n. F n)" "\n. I n (F n)" - using natural_partitionE by blast - let ?F = "{F n |n. True}" - have "I' (\S\?F. S)" - using assms(2) - proof (rule stable_by_infinite_unionE[of I']) - fix S assume "S \ {F n |n. True}" - then show "I' S" - using \\n. I n (F n)\ assms(1) by force - qed - moreover have "(\S\?F. S) = S" - using \S = (\n. F n)\ by auto - ultimately show "I' S" by blast - qed -qed - -corollary simpler_rule_while: - assumes "hyper_hoare_triple I C I" - and "stable_by_infinite_union I" - shows "hyper_hoare_triple I (While C) I" -proof - - let ?I = "\n. I" - have "hyper_hoare_triple (?I 0) (While C) (natural_partition ?I)" - using while_rule[of ?I C] - by (simp add: assms(1) assms(2) stable_by_union_and_constant_then_I) - then show ?thesis - by (simp add: assms(2) stable_by_union_and_constant_then_I) -qed - - text \Theorem 1\ theorem soundness: assumes "\ {A} C {B}" shows "\ {A} C {B}" using assms proof (induct rule: syntactic_HHT.induct) case (RuleSkip P) then show ?case using skip_rule by auto next case (RuleCons P P' Q' Q C) then show ?case using consequence_rule by blast next case (RuleExistsSet P C Q) then show ?case using rule_exists by blast next case (RuleSeq P C1 R C2 Q) then show ?case using seq_rule by meson next case (RuleIf P C1 Q1 C2 Q2) then show ?case using if_rule by blast next case (RuleAssume P b) then show ?case by (simp add: assume_rule) next case (RuleWhile I C) then show ?case using while_rule by blast next case (RuleAssign x e) then show ?case by (simp add: assign_rule) next case (RuleHavoc x) then show ?case using havoc_rule by fastforce qed subsection \Completeness\ definition complete where "complete P C Q \ (\ {P} C {Q} \ \ {P} C {Q})" lemma completeI: assumes "\ {P} C {Q} \ \ {P} C {Q}" shows "complete P C Q" by (simp add: assms complete_def) lemma completeE: assumes "complete P C Q" and "\ {P} C {Q}" shows "\ {P} C {Q}" using assms complete_def by auto lemma complete_if_aux: assumes "hyper_hoare_triple A (If C1 C2) B" shows "entails (\S'. \S. A S \ S' = sem C1 S \ sem C2 S) B" proof (rule entailsI) fix S' assume "\S. A S \ S' = sem C1 S \ sem C2 S" then show "B S'" by (metis assms hyper_hoare_tripleE sem_if) qed lemma complete_if: fixes P Q :: "('lvar, 'lval, 'pvar, 'pval) state hyperassertion" assumes "\P1 Q1 :: ('lvar, 'lval, 'pvar, 'pval) state hyperassertion. complete P1 C1 Q1" and "\P2 Q2 :: ('lvar, 'lval, 'pvar, 'pval) state hyperassertion. complete P2 C2 Q2" shows "complete P (If C1 C2) Q" proof (rule completeI) assume asm0: "\ {P} If C1 C2 {Q}" show "\ {P} stmt.If C1 C2 {Q}" proof (rule RuleCons) show "\ {exists (\V S. P S \ S = V)} stmt.If C1 C2 {exists (\V. join (\S. S = sem C1 V \ P V) (\S. S = sem C2 V))}" proof (rule RuleExistsSet) fix V show "\ {(\S. P S \ S = V)} stmt.If C1 C2 {join (\S. S = sem C1 V \ P V) (\S. S = sem C2 V)}" proof (rule RuleIf) show "\ {(\S. P S \ S = V)} C1 {\S. S = sem C1 V \ P V}" by (simp add: assms(1) completeE hyper_hoare_triple_def) show "\ {(\S. P S \ S = V)} C2 {\S. S = sem C2 V}" by (simp add: assms(2) completeE hyper_hoare_triple_def) qed qed show "entails P (exists (\V S. P S \ S = V))" by (simp add: entailsI exists_def) show "entails (exists (\V. join (\S. S = sem C1 V \ P V) (\S. S = sem C2 V))) Q" proof (rule entailsI) fix S assume "exists (\V. join (\S. S = sem C1 V \ P V) (\S. S = sem C2 V)) S" then obtain V where "join (\S. S = sem C1 V \ P V) (\S. S = sem C2 V) S" by (meson exists_def) then obtain S1 S2 where "S = S1 \ S2" "S1 = sem C1 V \ P V" "S2 = sem C2 V" by (simp add: join_def) then show "Q S" by (metis asm0 hyper_hoare_tripleE sem_if) qed qed qed lemma complete_seq_aux: assumes "hyper_hoare_triple A (Seq C1 C2) B" shows "\R. hyper_hoare_triple A C1 R \ hyper_hoare_triple R C2 B" proof - let ?R = "\S. \S'. A S' \ S = sem C1 S'" have "hyper_hoare_triple A C1 ?R" using hyper_hoare_triple_def by blast moreover have "hyper_hoare_triple ?R C2 B" proof (rule hyper_hoare_tripleI) fix S assume "\S'. A S' \ S = sem C1 S'" then obtain S' where asm0: "A S'" "S = sem C1 S'" by blast then show "B (sem C2 S)" by (metis assms hyper_hoare_tripleE sem_seq) qed ultimately show ?thesis by blast qed lemma complete_assume: "complete P (Assume b) Q" proof (rule completeI) assume asm0: "\ {P} Assume b {Q}" show "\ {P} Assume b {Q}" proof (rule RuleCons) show "\ { (\S. Q (Set.filter (b \ snd) S)) } (Assume b) {Q}" by (simp add: RuleAssume) show "entails P (\S. Q (Set.filter (b \ snd) S))" by (metis (mono_tags, lifting) asm0 assume_sem entails_def hyper_hoare_tripleE) show "entails Q Q" by (simp add: entailsI) qed qed lemma complete_skip: "complete P Skip Q" using completeI RuleSkip by (metis (mono_tags, lifting) entails_def hyper_hoare_triple_def sem_skip RuleCons) lemma complete_assign: "complete P (Assign x e) Q" proof (rule completeI) assume asm0: "\ {P} Assign x e {Q}" show "\ {P} Assign x e {Q}" proof (rule RuleCons) show "\ {(\S. Q {(l, \(x := e \)) |l \. (l, \) \ S})} Assign x e {Q}" by (simp add: RuleAssign) show "entails P (\S. Q {(l, \(x := e \)) |l \. (l, \) \ S})" proof (rule entailsI) fix S assume "P S" then show "Q {(l, \(x := e \)) |l \. (l, \) \ S}" by (metis asm0 hyper_hoare_triple_def sem_assign) qed show "entails Q Q" by (simp add: entailsI) qed qed lemma complete_havoc: "complete P (Havoc x) Q" proof (rule completeI) assume asm0: "\ {P} Havoc x {Q}" show "\ {P} Havoc x {Q}" proof (rule RuleCons) show "\ { (\S. Q { (l, \(x := v)) |l \ v. (l, \) \ S }) } (Havoc x) {Q}" using RuleHavoc by fast show "entails P (\S. Q {(l, \(x := v)) |l \ v. (l, \) \ S})" proof (rule entailsI) fix S assume "P S" then show "Q {(l, \(x := v)) |l \ v. (l, \) \ S}" by (metis asm0 hyper_hoare_triple_def sem_havoc) qed show "entails Q Q" by (simp add: entailsI) qed qed lemma complete_seq: assumes "\R. complete P C1 R" and "\R. complete R C2 Q" shows "complete P (Seq C1 C2) Q" by (meson RuleSeq assms(1) assms(2) completeE completeI complete_seq_aux) fun construct_inv where "construct_inv P C 0 = P" | "construct_inv P C (Suc n) = (\S. (\S'. S = sem C S' \ construct_inv P C n S'))" lemma iterate_sem_ind: assumes "construct_inv P C n S'" shows "\S. P S \ S' = iterate_sem n C S" using assms by (induct n arbitrary: S') (auto) lemma complete_while_aux: assumes "hyper_hoare_triple (\S. P S \ S = V) (While C) Q" shows "entails (natural_partition (construct_inv (\S. P S \ S = V) C)) Q" proof (rule entailsI) fix S assume "natural_partition (construct_inv (\S. P S \ S = V) C) S" then obtain F where asm0: "S = (\n. F n)" "\n. construct_inv (\S. P S \ S = V) C n (F n)" using natural_partitionE by blast then have "P (F 0) \ F 0 = V" by (metis (mono_tags, lifting) construct_inv.simps(1)) then have "Q (\n. iterate_sem n C (F 0))" using assms hyper_hoare_triple_def[of "\S. P S \ S = V" "While C" Q] sem_while by metis moreover have "\n. F n = iterate_sem n C V" proof - fix n obtain S' where "P S' \ S' = V" "F n = iterate_sem n C S'" using asm0(2) iterate_sem_ind by blast then show "F n = iterate_sem n C V" by simp qed ultimately show "Q S" using asm0(1) by auto qed lemma complete_while: fixes P Q :: "('lvar, 'lval, 'pvar, 'pval) state hyperassertion" assumes "\P' Q' :: ('lvar, 'lval, 'pvar, 'pval) state hyperassertion. complete P' C Q'" shows "complete P (While C) Q" proof (rule completeI) assume asm0: "hyper_hoare_triple P (While C) Q" let ?I = "\V. construct_inv (\S. P S \ S = V) C" have r: "\V. syntactic_HHT (?I V 0) (While C) (natural_partition (?I V))" proof (rule RuleWhile) fix V n show "syntactic_HHT (construct_inv (\S. P S \ S = V) C n) C (construct_inv (\S. P S \ S = V) C (Suc n))" by (meson assms completeE construct_inv.simps(2) hyper_hoare_tripleI) qed show "syntactic_HHT P (While C) Q" proof (rule RuleCons) show "syntactic_HHT (exists (\V. ?I V 0)) (While C) (exists (\V. ((natural_partition (?I V)))))" using r by (rule RuleExistsSet) show "entails P (exists (\V. construct_inv (\S. P S \ S = V) C 0))" by (simp add: entailsI exists_def) show "entails (exists (\V. natural_partition (construct_inv (\S. P S \ S = V) C))) Q" proof (rule entailsI) fix S' assume "exists (\V. natural_partition (construct_inv (\S. P S \ S = V) C)) S'" then obtain V where "natural_partition (construct_inv (\S. P S \ S = V) C) S'" by (meson exists_def) moreover have "entails (natural_partition (construct_inv (\S. P S \ S = V) C)) Q" proof (rule complete_while_aux) show "hyper_hoare_triple (\S. P S \ S = V) (While C) Q" using asm0 hyper_hoare_triple_def[of "\S. P S \ S = V"] hyper_hoare_triple_def[of P "While C" Q] by auto qed ultimately show "Q S'" by (simp add: entails_def) qed qed qed text \Theorem 2\ theorem completeness: fixes P Q :: "('lvar, 'lval, 'pvar, 'pval) state hyperassertion" assumes "\ {P} C {Q}" shows "\ {P} C {Q}" using assms proof (induct C arbitrary: P Q) case (Assign x1 x2) then show ?case using completeE complete_assign by fast next case (Seq C1 C2) then show ?case using complete_def complete_seq by meson next case (If C1 C2) then show ?case using complete_def complete_if by meson next case Skip then show ?case using complete_def complete_skip by meson next case (Havoc x) then show ?case by (simp add: completeE complete_havoc) next case (Assume b) then show ?case by (simp add: completeE complete_assume) next case (While C) then show ?case using complete_def complete_while by blast qed + + subsection \Disproving Hyper-Triples\ definition sat where "sat P \ (\S. P S)" -text \Theorem 4\ +text \Theorem 5\ theorem disproving_triple: "\ \ {P} C {Q} \ (\P'. sat P' \ entails P' P \ \ {P'} C {\S. \ Q S})" (is "?A \ ?B") proof assume "\ \ {P} C {Q}" then obtain S where asm0: "P S" "\ Q (sem C S)" using hyper_hoare_triple_def by blast let ?P = "\S'. S = S'" have "entails ?P P" by (simp add: asm0(1) entails_def) moreover have "\ {?P} C {\S. \ Q S}" by (simp add: asm0(2) hyper_hoare_triple_def) moreover have "sat ?P" by (simp add: sat_def) ultimately show ?B by blast next assume "\P'. sat P' \ entails P' P \ \ {P'} C {\S. \ Q S}" then obtain P' where asm0: "sat P'" "entails P' P" "\ {P'} C {\S. \ Q S}" by blast then obtain S where "P' S" by (meson sat_def) then show ?A using asm0(2) asm0(3) entailsE hyper_hoare_tripleE by (metis (no_types, lifting)) qed definition differ_only_by where "differ_only_by a b x \ (\y. y \ x \ a y = b y)" lemma differ_only_byI: assumes "\y. y \ x \ a y = b y" shows "differ_only_by a b x" by (simp add: assms differ_only_by_def) lemma diff_by_update: "differ_only_by (a(x := v)) a x" by (simp add: differ_only_by_def) lemma diff_by_comm: "differ_only_by a b x \ differ_only_by b a x" by (metis (mono_tags, lifting) differ_only_by_def) lemma diff_by_trans: assumes "differ_only_by a b x" and "differ_only_by b c x" shows "differ_only_by a c x" by (metis assms(1) assms(2) differ_only_by_def) definition not_free_var_of where "not_free_var_of P x \ (\states states'. (\i. differ_only_by (fst (states i)) (fst (states' i)) x \ snd (states i) = snd (states' i)) \ (states \ P \ states' \ P))" lemma not_free_var_ofE: assumes "not_free_var_of P x" and "\i. differ_only_by (fst (states i)) (fst (states' i)) x" and "\i. snd (states i) = snd (states' i)" and "states \ P" shows "states' \ P" using not_free_var_of_def[of P x] assms by blast subsection \Synchronized Rule for Branching\ definition combine where "combine from_nat x P1 P2 S \ P1 (Set.filter (\\. fst \ x = from_nat 1) S) \ P2 (Set.filter (\\. fst \ x = from_nat 2) S)" lemma combineI: assumes "P1 (Set.filter (\\. fst \ x = from_nat 1) S) \ P2 (Set.filter (\\. fst \ x = from_nat 2) S)" shows "combine from_nat x P1 P2 S" by (simp add: assms combine_def) definition modify_lvar_to where "modify_lvar_to x v \ = ((fst \)(x := v), snd \)" lemma logical_var_in_sem_same: assumes "\\. \ \ S \ fst \ x = a" and "\' \ sem C S" shows "fst \' x = a" by (metis assms(1) assms(2) fst_conv in_sem) lemma recover_after_sem: assumes "a \ b" and "\\. \ \ S1 \ fst \ x = a" and "\\. \ \ S2 \ fst \ x = b" shows "sem C S1 = Set.filter (\\. fst \ x = a) (sem C (S1 \ S2))" (is "?A = ?B") proof have r: "sem C (S1 \ S2) = sem C S1 \ sem C S2" by (simp add: sem_union) moreover have r1: "\\'. \' \ sem C S1 \ fst \' x = a" by (metis assms(2) fst_conv in_sem) moreover have r2: "\\'. \' \ sem C S2 \ fst \' x = b" by (metis assms(3) fst_conv in_sem) show "?B \ ?A" proof (rule subsetPairI) fix l \ assume "(l, \) \ Set.filter (\\. fst \ x = a) (sem C (S1 \ S2))" then show "(l, \) \ sem C S1" using assms(1) r r2 by auto qed show "?A \ ?B" by (simp add: r r1 subsetI) qed lemma injective_then_ok: assumes "a \ b" and "S1' = (modify_lvar_to x a) ` S1" and "S2' = (modify_lvar_to x b) ` S2" shows "Set.filter (\\. fst \ x = a) (S1' \ S2') = S1'" (is "?A = ?B") proof show "?B \ ?A" proof (rule subsetI) fix y assume "y \ S1'" then have "fst y x = a" using modify_lvar_to_def assms(2) by (metis (mono_tags, lifting) fst_conv fun_upd_same image_iff) then show "y \ Set.filter (\\. fst \ x = a) (S1' \ S2')" by (simp add: \y \ S1'\) qed show "?A \ ?B" proof fix y assume "y \ ?A" then have "y \ S2'" by (metis (mono_tags, lifting) assms(1) assms(3) fun_upd_same image_iff member_filter modify_lvar_to_def prod.sel(1)) then show "y \ ?B" using \y \ Set.filter (\\. fst \ x = a) (S1' \ S2')\ by auto qed qed definition not_free_var_hyper where "not_free_var_hyper x P \ (\S v. P S \ P ((modify_lvar_to x v) ` S))" definition injective where "injective f \ (\a b. a \ b \ f a \ f b)" lemma sem_of_modify_lvar: "sem C ((modify_lvar_to r v) ` S) = (modify_lvar_to r v) ` (sem C S)" (is "?A = ?B") proof show "?A \ ?B" proof (rule subsetI) fix y assume asm0: "y \ ?A" then obtain x where "x \ (modify_lvar_to r v) ` S" "single_sem C (snd x) (snd y)" "fst x = fst y" by (metis fst_conv in_sem snd_conv) then obtain xx where "xx \ S" "x = modify_lvar_to r v xx" by blast then have "(fst xx, snd y) \ sem C S" by (metis \\C, snd x\ \ snd y\ fst_conv in_sem modify_lvar_to_def prod.collapse snd_conv) then show "y \ ?B" by (metis \fst x = fst y\ \x = modify_lvar_to r v xx\ fst_eqD modify_lvar_to_def prod.exhaust_sel rev_image_eqI snd_eqD) qed show "?B \ ?A" proof (rule subsetI) fix y assume "y \ modify_lvar_to r v ` sem C S" then obtain yy where "y = modify_lvar_to r v yy" "yy \ sem C S" by blast then obtain x where "x \ S" "fst x = fst yy" "single_sem C (snd x) (snd yy)" by (metis fst_conv in_sem snd_conv) then have "fst (modify_lvar_to r v x) = fst y" by (simp add: \y = modify_lvar_to r v yy\ modify_lvar_to_def) then show "y \ sem C (modify_lvar_to r v ` S)" by (metis (mono_tags, lifting) \\C, snd x\ \ snd yy\ \x \ S\ \y = modify_lvar_to r v yy\ fst_conv image_eqI in_sem modify_lvar_to_def snd_conv) qed qed -text \Proposition 15 (appendix C).\ - -theorem if_sync_rule: - assumes "\ {P} C1 {P1}" - and "\ {P} C2 {P2}" - and "\ {combine from_nat x P1 P2} C {combine from_nat x R1 R2}" - and "\ {R1} C1' {Q1}" - and "\ {R2} C2' {Q2}" - - and "not_free_var_hyper x P1" - and "not_free_var_hyper x P2" - and "injective (from_nat :: nat \ 'a)" - - and "not_free_var_hyper x R1" - and "not_free_var_hyper x R2" - - shows "\ {P} If (Seq C1 (Seq C C1')) (Seq C2 (Seq C C2')) {join Q1 Q2}" -proof (rule hyper_hoare_tripleI) - fix S assume asm0: "P S" - have r0: "sem (stmt.If (Seq C1 (Seq C C1')) (Seq C2 (Seq C C2'))) S - = sem C1' (sem C (sem C1 S)) \ sem C2' (sem C (sem C2 S))" - by (simp add: sem_if sem_seq) - moreover have "P1 (sem C1 S) \ P2 (sem C2 S)" - using asm0 assms(1) assms(2) hyper_hoare_tripleE by blast - - let ?S1 = "(modify_lvar_to x (from_nat 1)) ` (sem C1 S)" - let ?S2 = "(modify_lvar_to x (from_nat 2)) ` (sem C2 S)" - let ?f1 = "Set.filter (\\. fst \ x = from_nat 1)" - let ?f2 = "Set.filter (\\. fst \ x = from_nat 2)" - - have r: "from_nat 1 \ from_nat 2" - by (metis Suc_1 assms(8) injective_def n_not_Suc_n) - - have "P1 ?S1 \ P2 ?S2" - by (meson \P1 (sem C1 S) \ P2 (sem C2 S)\ assms(6) assms(7) not_free_var_hyper_def) - moreover have rr1: "Set.filter (\\. fst \ x = from_nat 1) (?S1 \ ?S2) = ?S1" - using injective_then_ok[of "from_nat 1" "from_nat 2" ?S1 x] - by (metis (no_types, lifting) assms(8) injective_def num.simps(4) one_eq_numeral_iff) - moreover have rr2: "Set.filter (\\. fst \ x = from_nat 2) (?S1 \ ?S2) = ?S2" - using injective_then_ok[of "from_nat 2" "from_nat 1" ?S2 x] - by (metis (no_types, lifting) assms(8) injective_def one_eq_numeral_iff sup_commute verit_eq_simplify(10)) - ultimately have "combine from_nat x P1 P2 (?S1 \ ?S2)" - by (metis combineI) - then have "combine from_nat x R1 R2 (sem C (?S1 \ ?S2))" - using assms(3) hyper_hoare_tripleE by blast - moreover have "?f1 (sem C (?S1 \ ?S2)) = sem C ?S1" - using recover_after_sem[of "from_nat 1" "from_nat 2" ?S1 x ?S2] r rr1 rr2 - member_filter[of _ "\\. fst \ x = from_nat 1"] member_filter[of _ "\\. fst \ x = from_nat 2"] - by metis - then have "R1 (sem C ?S1)" - by (metis (mono_tags) calculation combine_def) - then have "R1 (sem C (sem C1 S))" - by (metis assms(9) not_free_var_hyper_def sem_of_modify_lvar) - moreover have "?f2 (sem C (?S1 \ ?S2)) = sem C ?S2" - using recover_after_sem[of "from_nat 2" "from_nat 1" ?S2 x ?S1] r rr1 rr2 sup_commute[of ] - member_filter[of _ "\\. fst \ x = from_nat 1" "?S1 \ ?S2"] member_filter[of _ "\\. fst \ x = from_nat 2" "?S1 \ ?S2"] - by metis - then have "R2 (sem C ?S2)" - by (metis (mono_tags) calculation(1) combine_def) - then have "R2 (sem C (sem C2 S))" - by (metis assms(10) not_free_var_hyper_def sem_of_modify_lvar) - - then show "join Q1 Q2 (sem (stmt.If (Seq C1 (Seq C C1')) (Seq C2 (Seq C C2'))) S)" - by (metis (full_types) r0 assms(4) assms(5) calculation(2) hyper_hoare_tripleE join_def) -qed - end \ No newline at end of file diff --git a/thys/HyperHoareLogic/Loops.thy b/thys/HyperHoareLogic/Loops.thy new file mode 100644 --- /dev/null +++ b/thys/HyperHoareLogic/Loops.thy @@ -0,0 +1,1029 @@ +theory Loops + imports Logic HOL.Wellfounded Expressivity +begin + +section \Rules for Loops\ + +definition lnot where + "lnot b \ = (\b \)" + +definition if_then_else where + "if_then_else b C1 C2 = If (Assume b;; C1) (Assume (lnot b);; C2)" + +definition low_exp where + "low_exp e S = (\\ \'. \ \ S \ \' \ S \ (e (snd \) = e (snd \')))" + +lemma low_exp_lnot: + "low_exp b S \ low_exp (lnot b) S" + by (simp add: lnot_def low_exp_def) + +definition holds_forall where + "holds_forall b S \ (\\\S. b (snd \))" + +lemma holds_forallI: + assumes "\\. \ \ S \ b (snd \)" + shows "holds_forall b S" + using assms holds_forall_def by blast + +lemma low_exp_two_cases: + assumes "low_exp b S" + shows "holds_forall b S \ holds_forall (lnot b) S" + by (metis assms holds_forall_def lnot_def low_exp_def) + +lemma sem_assume_low_exp: + assumes "holds_forall b S" + shows "sem (Assume b) S = S" + and "sem (Assume (lnot b)) S = {}" + using assume_sem[of b S] assms holds_forall_def[of b S] apply fastforce + using assume_sem[of "lnot b" S] assms holds_forall_def[of b S] lnot_def[of b] + by fastforce + +lemma sem_assume_low_exp_seq: + assumes "holds_forall b S" + shows "sem (Assume b;; C) S = sem C S" + and "sem (Assume (lnot b);; C) S = {}" + apply (simp add: assms sem_assume_low_exp(1) sem_seq) + by (metis assms empty_iff equals0I in_sem sem_assume_low_exp(2) sem_seq) + +lemma lnot_involution: + "lnot (lnot b) = b" +proof (rule ext) + fix x show "lnot (lnot b) x = b x" + by (simp add: lnot_def) +qed + +lemma sem_if_then_else: + shows "holds_forall b S \ sem (if_then_else b C1 C2) S = sem C1 S" + and "holds_forall (lnot b) S \ sem (if_then_else b C1 C2) S = sem C2 S" + apply (simp add: if_then_else_def sem_assume_low_exp_seq(1) sem_assume_low_exp_seq(2) sem_if) + by (metis (no_types, opaque_lifting) if_then_else_def lnot_involution sem_assume_low_exp_seq(1) sem_assume_low_exp_seq(2) sem_if sup_bot_left) + +lemma if_synchronized_aux: + assumes "\ {P} C1 {Q}" + and "\ {P} C2 {Q}" + and "entails P (low_exp b)" + shows "\ {P} if_then_else b C1 C2 {Q}" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "P S" + then have r: "low_exp b S" using assms(3) entailsE + by metis + show "Q (sem (if_then_else b C1 C2) S)" + proof (cases "holds_forall b S") + case True + then show ?thesis + by (metis asm0 assms(1) hyper_hoare_triple_def sem_if_then_else(1)) + next + case False + then show ?thesis + by (metis asm0 assms(2) hyper_hoare_tripleE low_exp_two_cases r sem_if_then_else(2)) + qed +qed + +theorem if_synchronized: + assumes "\ {conj P (holds_forall b)} C1 {Q}" + and "\ {conj P (holds_forall (lnot b))} C2 {Q}" + shows "\ {conj P (low_exp b)} if_then_else b C1 C2 {Q}" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "conj P (low_exp b) S" + show "Q (sem (if_then_else b C1 C2) S)" + proof (cases "holds_forall b S") + case True + then show ?thesis + by (metis asm0 assms(1) conj_def hyper_hoare_triple_def sem_if_then_else(1)) + next + case False + then show ?thesis + by (metis asm0 assms(2) conj_def hyper_hoare_triple_def low_exp_two_cases sem_if_then_else(2)) + qed +qed + + +definition while_cond where + "while_cond b C = While (Assume b;; C);; Assume (lnot b)" + + +lemma while_synchronized_rec: + assumes "\n. \ {conj (I n) (holds_forall b)} Assume b;; C {conj (I (Suc n)) (low_exp b)}" + and "conj (I 0) (low_exp b) S" + shows "conj (I n) (low_exp b) (iterate_sem n (Assume b;; C) S) \ holds_forall (lnot b) (iterate_sem n (Assume b;; C) S)" + using assms +proof (induct n) + case (Suc n) + then have r: "conj (I n) (low_exp b) (iterate_sem n (Assume b;; C) S) \ holds_forall (lnot b) (iterate_sem n (Assume b;; C) S)" + by blast + then show ?case + proof (cases "conj (I n) (holds_forall b) (iterate_sem n (Assume b;; C) S)") + case True + then show ?thesis + using Suc.prems(1) hyper_hoare_tripleE by fastforce + next + case False + then have "holds_forall (lnot b) (iterate_sem n (Assume b;; C) S)" + by (metis conj_def low_exp_two_cases r) + then have "iterate_sem (Suc n) (Assume b;; C) S = {}" + by (metis iterate_sem.simps(2) lnot_involution sem_assume_low_exp_seq(2)) + then show ?thesis + by (simp add: holds_forall_def) + qed +qed (auto) + +lemma false_then_empty_later: + assumes "holds_forall (lnot b) (iterate_sem n (Assume b;; C) S)" + and "m > n" + shows "iterate_sem m (Assume b;; C) S = {}" + using assms +proof (induct "m - n" arbitrary: n m) + case (Suc x) + then show ?case + proof (cases x) + case 0 + then show ?thesis + by (metis One_nat_def Suc.hyps(2) Suc.prems(1) Suc.prems(2) Suc_eq_plus1 iterate_sem.simps(2) le_add_diff_inverse linorder_not_le lnot_involution order.asym sem_assume_low_exp_seq(2)) + next + case (Suc nat) + then have "m - 1 > n" + using Suc.hyps(2) by auto + then have "iterate_sem (m-1) (Assume b ;; C) S = {}" + by (metis (no_types, lifting) Suc.hyps(1) Suc.hyps(2) Suc.prems(1) diff_Suc_1 diff_commute) + then show ?thesis + by (metis Nat.lessE Suc.prems(1) Suc.prems(2) diff_Suc_1 iterate_sem.simps(2) sem_assume_low_exp_seq(2) sem_seq) + qed +qed (simp) + +lemma split_union_triple: + "(\(m::nat). f m) = (\m\{m |m. m < n}. f m) \ f n \ (\m\{m |m. m > n}. f m)" (is "?A = ?B") +proof + show "?B \ ?A" + by blast + show "?A \ ?B" + proof + fix x assume "x \ ?A" + then obtain m where "x \ f m" + by blast + then have "m < n \ m = n \ m > n" + by force + then show "x \ ?B" + using \x \ f m\ by auto + qed +qed + + +lemma sem_union_swap: + "sem C (\x\S. f x) = (\x\S. sem C (f x))" (is "?A = ?B") +proof + show "?A \ ?B" + proof + fix y assume "y \ ?A" + then obtain x where "x \ S" "y \ sem C (f x)" + using UN_iff in_sem[of y C] by force + then show "y \ ?B" + by blast + qed + show "?B \ ?A" + by (simp add: SUP_least SUP_upper sem_monotonic) +qed + + + +lemma while_synchronized_case_1: + assumes "\m. m < n \ holds_forall b (iterate_sem m (Assume b;; C) S)" + and "holds_forall (lnot b) (iterate_sem n (Assume b;; C) S)" + and "\n. \ {conj (I n) (holds_forall b)} Assume b;; C {conj (I (Suc n)) (low_exp b)}" + and "conj (I 0) (low_exp b) S" + shows "sem (while_cond b C) S = iterate_sem n (Assume b;; C) S" +proof - + have "\m. m > n \ iterate_sem m (Assume b;; C) S = {}" + using assms(2) false_then_empty_later by blast + moreover have "sem (While (Assume b;; C)) S = + (\m\{m|m. m < n}. iterate_sem m (Assume b ;; C) S) \ iterate_sem n (Assume b ;; C) S \ (\m\{m|m. m > n}. iterate_sem m (Assume b ;; C) S)" + using sem_while[of "Assume b;; C" S] split_union_triple by metis + ultimately have "sem (While (Assume b;; C)) S = (\m\{m|m. m < n}. iterate_sem m (Assume b ;; C) S) \ iterate_sem n (Assume b ;; C) S " + by auto + moreover have "\m. m < n \ sem (Assume (lnot b)) (iterate_sem m (Assume b ;; C) S) = {}" + using assms(1) sem_assume_low_exp(2) by blast + then have "sem (Assume (lnot b)) (\m\{m|m. m < n}. iterate_sem m (Assume b ;; C) S) = {}" + by (simp add: sem_union_swap) + then have "sem (while_cond b C) S = sem (Assume (lnot b)) (iterate_sem n (Assume b ;; C) S)" + by (simp add: calculation sem_seq sem_union while_cond_def) + then show ?thesis + using assms(2) sem_assume_low_exp(1) by blast +qed + +lemma while_synchronized_case_2: + assumes "\m. holds_forall b (iterate_sem m (Assume b;; C) S)" + and "\n. \ {conj (I n) (holds_forall b)} Assume b;; C {conj (I (Suc n)) (low_exp b)}" + and "conj (I 0) (low_exp b) S" + shows "sem (while_cond b C) S = {}" +proof - + have "sem (While (Assume b ;; C)) S = (\n. iterate_sem n (Assume b ;; C) S)" + by (simp add: sem_while) + then have "holds_forall b (sem (While (Assume b;; C)) S)" + by (metis (no_types, lifting) UN_iff assms(1) holds_forall_def) + then show ?thesis + by (simp add: sem_assume_low_exp(2) sem_seq while_cond_def) +qed + +definition emp where + "emp S \ S = {}" + +lemma holds_forall_empty: + "holds_forall b {}" + by (simp add: holds_forall_def) + +definition exists where + "exists I S \ (\n. I n S)" + +theorem while_synchronized: + assumes "\n. \ {conj (I n) (holds_forall b)} C {conj (I (Suc n)) (low_exp b)}" + shows "\ {conj (I 0) (low_exp b)} while_cond b C {conj (disj (exists I) emp) (holds_forall (lnot b))}" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "conj (I 0) (low_exp b) S" + have triple: "\n. \ {conj (I n) (holds_forall b)} Assume b ;; C {conj (I (Suc n)) (low_exp b)}" + proof (rule hyper_hoare_tripleI) + fix n S assume "conj (I n) (holds_forall b) S" + then have "sem (Assume b) S = S" + by (simp add: conj_def sem_assume_low_exp(1)) + then show "conj (I (Suc n)) (low_exp b) (sem (Assume b ;; C) S)" + by (metis \conj (I n) (holds_forall b) S\ assms hyper_hoare_tripleE sem_seq) + qed + show "conj (disj (exists I) emp) (holds_forall (lnot b)) (sem (while_cond b C) S)" + proof (cases "\m. holds_forall b (iterate_sem m (Assume b;; C) S)") + case True + then have "sem (while_cond b C) S = {}" + using while_synchronized_case_2[of b C S I] + by (metis (no_types, opaque_lifting) asm0 assms hyper_hoare_tripleI conj_def sem_assume_low_exp(1) seq_rule) + then show ?thesis + by (simp add: disj_def conj_def emp_def holds_forall_empty) + next + case False + then have F: "\ (\m. holds_forall b (iterate_sem m (Assume b;; C) S))" by simp + have "\n. (\m. m < n \ holds_forall b (iterate_sem m (Assume b;; C) S)) \ holds_forall (lnot b) (iterate_sem n (Assume b;; C) S)" + proof (cases "\n. \ holds_forall b (iterate_sem n (Assume b;; C) S) \ iterate_sem n (Assume b;; C) S \ {}") + case True + then obtain n where "\ holds_forall b (iterate_sem n (Assume b;; C) S) \ iterate_sem n (Assume b;; C) S \ {}" + by blast + then have "holds_forall (lnot b) (iterate_sem n (Assume b;; C) S)" + by (metis asm0 conj_def low_exp_two_cases triple while_synchronized_rec) + moreover have "\m. m < n \ holds_forall b (iterate_sem m (Assume b;; C) S)" + proof - + fix m assume asm1: "m < n" + show "holds_forall b (iterate_sem m (Assume b;; C) S)" + proof (rule ccontr) + assume "\ holds_forall b (iterate_sem m (Assume b ;; C) S)" + then have "holds_forall (lnot b) (iterate_sem m (Assume b;; C) S)" + by (metis asm0 conj_def low_exp_two_cases triple while_synchronized_rec) + then have "iterate_sem n (Assume b;; C) S = {}" + using asm1 false_then_empty_later by blast + then show False + using \\ holds_forall b (iterate_sem n (Assume b ;; C) S) \ iterate_sem n (Assume b ;; C) S \ {}\ by fastforce + qed + qed + ultimately show ?thesis + by blast + next + case False + then have "\n. holds_forall b (iterate_sem n (Assume b;; C) S)" + using holds_forall_empty by fastforce + then show ?thesis using F by blast + qed + then obtain n where "\m. m < n \ holds_forall b (iterate_sem m (Assume b;; C) S)" + and "holds_forall (lnot b) (iterate_sem n (Assume b;; C) S)" + by blast + then have "sem (while_cond b C) S = iterate_sem n (Assume b;; C) S" + using triple + proof (rule while_synchronized_case_1) + qed (simp_all add: asm0) + moreover have "I n (iterate_sem n (Assume b;; C) S)" + proof (cases n) + case 0 + then show ?thesis + by (metis asm0 iterate_sem.simps(1) conj_def) + next + case (Suc k) + then have "conj (I k) (low_exp b) (iterate_sem k (Assume b ;; C) S) \ holds_forall (lnot b) (iterate_sem k (Assume b ;; C) S)" + using while_synchronized_rec[of I b C S k] asm0 triple by blast + then show ?thesis + proof (cases "conj (I k) (low_exp b) (iterate_sem k (Assume b ;; C) S)") + case True + then show ?thesis + using conj_def[of _ "holds_forall b"] conj_def[of _ "low_exp b"] Suc + \\m. m < n \ holds_forall b (iterate_sem m (Assume b ;; C) S)\ assms + hyper_hoare_triple_def[of ] iterate_sem.simps(2) lessI sem_assume_low_exp(1)[of b "iterate_sem k (Assume b ;; C) S"] + sem_seq[of "Assume b" C] by metis + next + case False + then show ?thesis + by (metis F Suc \\m. m < n \ holds_forall b (iterate_sem m (Assume b ;; C) S)\ \conj (I k) (low_exp b) (iterate_sem k (Assume b ;; C) S) \ holds_forall (lnot b) (iterate_sem k (Assume b ;; C) S)\ empty_iff false_then_empty_later holds_forall_def not_less_eq) + qed + qed + ultimately show ?thesis + by (metis disj_def Loops.exists_def \holds_forall (lnot b) (iterate_sem n (Assume b ;; C) S)\ conj_def) + qed +qed + +lemma WhileSync_simpler: + assumes "\ {conj I (holds_forall b)} C {conj I (low_exp b)}" + shows "\ {conj I (low_exp b)} while_cond b C {conj (disj I emp) (holds_forall (lnot b))}" + using assms while_synchronized[of "\n. I"] + by (simp add: disj_def Loops.exists_def conj_def hyper_hoare_triple_def) + +definition if_then where + "if_then b C = If (Assume b;; C) (Assume (lnot b))" + +definition filter_exp where + "filter_exp b S = Set.filter (b \ snd) S" + +lemma filter_exp_union: + "filter_exp b (S1 \ S2) = filter_exp b S1 \ filter_exp b S2" (is "?A = ?B") +proof + show "?A \ ?B" + proof + fix x assume "x \ ?A" + then show "x \ ?B" + proof (cases "x \ S1") + case True + then show ?thesis + by (metis UnI1 \x \ filter_exp b (S1 \ S2)\ filter_exp_def member_filter) + next + case False + then show ?thesis + by (metis Un_iff \x \ filter_exp b (S1 \ S2)\ filter_exp_def member_filter) + qed + qed + show "?B \ ?A" + by (simp add: filter_exp_def subset_iff) +qed + +lemma filter_exp_union_general: + "filter_exp b (\x. f x) = (\x. filter_exp b (f x))" (is "?A = ?B") +proof + show "?A \ ?B" + proof + fix y assume "y \ ?A" + then obtain x where "y \ f x" + by (metis UN_iff filter_exp_def member_filter) + then show "y \ ?B" + by (metis UN_iff \y \ filter_exp b (\ (range f))\ filter_exp_def member_filter) + qed + show "?B \ ?A" + by (simp add: filter_exp_def subset_iff) +qed + +lemma filter_exp_contradict: + "filter_exp b (filter_exp (lnot b) S) = {}" + by (metis (mono_tags, lifting) all_not_in_conv filter_exp_def lnot_def member_filter o_apply) + +lemma filter_exp_same: + "filter_exp b (filter_exp b S) = filter_exp b S" (is "?A = ?B") +proof + show "?A \ ?B" + by (metis filter_exp_def member_filter subsetI) + show "?B \ ?A" + by (metis filter_exp_def member_filter subrelI) +qed + +lemma if_then_sem: + "sem (if_then b C) S = sem C (filter_exp b S) \ filter_exp (lnot b) S" + by (simp add: assume_sem filter_exp_def if_then_def sem_if sem_seq) + +fun union_up_to_n where + "union_up_to_n C S 0 = iterate_sem 0 C S" +| "union_up_to_n C S (Suc n) = iterate_sem (Suc n) C S \ union_up_to_n C S n" + +lemma union_up_to_increasing: + assumes "m \ n" + shows "union_up_to_n C S m \ union_up_to_n C S n" + using assms +proof (induct "n - m" arbitrary: m n) + case (Suc x) + then show ?case + by (simp add: lift_Suc_mono_le) +qed (simp) + +lemma union_union_up_to_n_equiv_aux: + "union_up_to_n C S n \ (\m. iterate_sem m C S)" +proof (induct n) + case 0 + then show ?case + by (metis UN_upper iso_tuple_UNIV_I union_up_to_n.simps(1)) +next + case (Suc n) + show ?case + proof + fix x assume "x \ union_up_to_n C S (Suc n)" (* \ (\m. iterate_sem m C S) *) + then have "x \ iterate_sem (Suc n) C S \ x \ union_up_to_n C S n" + by simp + then show "x \ (\m. iterate_sem m C S)" + using Suc by blast + qed +qed + +lemma union_union_up_to_n_equiv: + "(\n. union_up_to_n C S n) = (\n. iterate_sem n C S)" (is "?A = ?B") +proof + show "?B \ ?A" + by (metis (no_types, lifting) SUP_subset_mono UnCI subsetI union_up_to_n.elims) + show "?A \ ?B" + by (simp add: SUP_le_iff union_union_up_to_n_equiv_aux) +qed + + +lemma filter_exp_union_itself: + "filter_exp b S \ S = S" + by (metis Un_absorb1 filter_exp_def member_filter subsetI) + +lemma iterate_sem_equiv: + "iterate_sem m (if_then b C) S + = filter_exp (lnot b) (union_up_to_n (Assume b;; C) S m) \ iterate_sem m (Assume b;; C) S" +proof (induct m) + case 0 + have "union_up_to_n (Assume b ;; C) S 0 = S" + by auto + then show "iterate_sem 0 (if_then b C) S = filter_exp (lnot b) (union_up_to_n (Assume b ;; C) S 0) \ iterate_sem 0 (Assume b ;; C) S" + using filter_exp_def + by (metis Un_commute iterate_sem.simps(1) member_filter subset_iff sup.order_iff) +next + case (Suc m) + + let ?S = "iterate_sem m (if_then b C) S" + let ?SU = "union_up_to_n (Assume b ;; C) S m" + let ?SN = "iterate_sem m (Assume b ;; C) S" + have "iterate_sem (Suc m) (if_then b C) S = sem C (filter_exp b ?S) \ filter_exp (lnot b) ?S" + by (simp add: if_then_sem) + also have "... = sem C (filter_exp b (filter_exp (lnot b) ?SU)) \ sem C (filter_exp b ?SN) + \ filter_exp (lnot b) (filter_exp (lnot b) ?SU) \ filter_exp (lnot b) ?SN" + by (simp add: Suc filter_exp_union sem_union sup_assoc) + also have "... = sem C (filter_exp b ?SN) \ filter_exp (lnot b) ?SU \ filter_exp (lnot b) ?SN" + by (metis Un_empty_left filter_exp_contradict filter_exp_same sem_union) + moreover have "iterate_sem (Suc m) (Assume b ;; C) S = sem C (filter_exp b ?SN)" + by (simp add: assume_sem filter_exp_def sem_seq) + moreover have "union_up_to_n (Assume b ;; C) S (Suc m) = sem C (filter_exp b ?SN) \ ?SU" + using calculation(3) by force + moreover have "filter_exp (lnot b) (union_up_to_n (Assume b ;; C) S (Suc m)) \ iterate_sem (Suc m) (Assume b ;; C) S + = filter_exp (lnot b) (sem C (filter_exp b ?SN) \ ?SU) \ sem C (filter_exp b ?SN)" + using calculation(3) by force + then have "... = filter_exp (lnot b) ?SU \ sem C (filter_exp b ?SN)" + using filter_exp_union_itself[of "lnot b"] filter_exp_union[of "lnot b"] Un_commute sup_assoc by blast + moreover have "?SN \ ?SU" + by (metis UnCI subsetI union_up_to_n.elims) + ultimately have "filter_exp (lnot b) ?SU \ sem C (filter_exp b ?SN) + = sem C (filter_exp b ?SN) \ filter_exp (lnot b) ?SU \ filter_exp (lnot b) ?SN" + using filter_exp_union[of "lnot b" ?SU ?SN] + using Un_commute[of "filter_exp (lnot b) ?SU" "sem C (filter_exp b ?SN)"] + sup.orderE sup_assoc[of "sem C (filter_exp b ?SN)"] by metis + then show ?case + using \filter_exp (lnot b) (sem C (filter_exp b (iterate_sem m (Assume b ;; C) S)) \ union_up_to_n (Assume b ;; C) S m) \ sem C (filter_exp b (iterate_sem m (Assume b ;; C) S)) = filter_exp (lnot b) (union_up_to_n (Assume b ;; C) S m) \ sem C (filter_exp b (iterate_sem m (Assume b ;; C) S))\ \iterate_sem (Suc m) (Assume b ;; C) S = sem C (filter_exp b (iterate_sem m (Assume b ;; C) S))\ \iterate_sem (Suc m) (if_then b C) S = sem C (filter_exp b (iterate_sem m (if_then b C) S)) \ filter_exp (lnot b) (iterate_sem m (if_then b C) S)\ \sem C (filter_exp b (filter_exp (lnot b) (union_up_to_n (Assume b ;; C) S m))) \ sem C (filter_exp b (iterate_sem m (Assume b ;; C) S)) \ filter_exp (lnot b) (filter_exp (lnot b) (union_up_to_n (Assume b ;; C) S m)) \ filter_exp (lnot b) (iterate_sem m (Assume b ;; C) S) = sem C (filter_exp b (iterate_sem m (Assume b ;; C) S)) \ filter_exp (lnot b) (union_up_to_n (Assume b ;; C) S m) \ filter_exp (lnot b) (iterate_sem m (Assume b ;; C) S)\ \sem C (filter_exp b (iterate_sem m (if_then b C) S)) \ filter_exp (lnot b) (iterate_sem m (if_then b C) S) = sem C (filter_exp b (filter_exp (lnot b) (union_up_to_n (Assume b ;; C) S m))) \ sem C (filter_exp b (iterate_sem m (Assume b ;; C) S)) \ filter_exp (lnot b) (filter_exp (lnot b) (union_up_to_n (Assume b ;; C) S m)) \ filter_exp (lnot b) (iterate_sem m (Assume b ;; C) S)\ by auto +qed + + +lemma sem_while_with_if: + "sem (while_cond b C) S = filter_exp (lnot b) (\n. iterate_sem n (if_then b C) S)" +proof - + have "(\n. iterate_sem n (if_then b C) S) + = (\n. filter_exp (lnot b) (union_up_to_n (Assume b;; C) S n) \ iterate_sem n (Assume b;; C) S)" + by (simp add: iterate_sem_equiv) + also have "... = filter_exp (lnot b) (\n. union_up_to_n (Assume b;; C) S n) \ (\n. iterate_sem n (Assume b;; C) S)" + by (simp add: complete_lattice_class.SUP_sup_distrib filter_exp_union_general) + also have "... = filter_exp (lnot b) (\n. iterate_sem n (Assume b;; C) S) \ (\n. iterate_sem n (Assume b;; C) S)" + by (simp add: union_union_up_to_n_equiv) + also have "... = (\n. iterate_sem n (Assume b;; C) S)" + by (meson filter_exp_union_itself) + moreover have "sem (while_cond b C) S = filter_exp (lnot b) (\n. iterate_sem n (Assume b ;; C) S)" + by (simp add: assume_sem filter_exp_def sem_seq sem_while while_cond_def) + ultimately show ?thesis + by presburger +qed + +lemma iterate_sem_assume_increasing: + "filter_exp (lnot b) (iterate_sem n (if_then b C) S) \ filter_exp (lnot b) (iterate_sem (Suc n) (if_then b C) S)" + by (metis (no_types, lifting) UnCI filter_exp_def if_then_sem iterate_sem.simps(2) member_filter subsetI) + +lemma iterate_sem_assume_increasing_union_up_to: + "filter_exp (lnot b) (iterate_sem n (if_then b C) S) = filter_exp (lnot b) (union_up_to_n (if_then b C) S n)" +proof (induct n) + case (Suc n) + then show ?case + by (metis filter_exp_union iterate_sem_assume_increasing sup.orderE union_up_to_n.simps(2)) +qed (simp) + +(* Set becomes larger *) +definition ascending :: "(nat \ 'b set) \ bool" where + "ascending S \ (\n m. n \ m \ S n \ S m)" + +lemma ascendingI_direct: + assumes "\n m. n \ m \ S n \ S m" + shows "ascending S" + by (simp add: ascending_def assms) + +lemma ascendingI: + assumes "\n. S n \ S (Suc n)" + shows "ascending S" +proof (rule ascendingI_direct) + fix n m :: nat assume asm0: "n \ m" + moreover have "n \ m \ S n \ S m" + proof (induct "m - n" arbitrary: m n) + case (Suc x) + then show ?case + using assms lift_Suc_mono_le by blast + qed (simp) + ultimately show "S n \ S m" + by blast +qed + + + +definition upwards_closed where + "upwards_closed P P_inf \ (\S. ascending S \ (\n. P n (S n)) \ P_inf (\n. S n))" + +lemma upwards_closedI: + assumes "\S. ascending S \ (\n. P n (S n)) \ P_inf (\n. S n)" + shows "upwards_closed P P_inf" + using assms upwards_closed_def by blast + +lemma upwards_closedE: + assumes "upwards_closed P P_inf" + and "ascending S" + and "\n. P n (S n)" + shows "P_inf (\n. S n)" + using assms(1) assms(2) assms(3) upwards_closed_def by blast + +lemma ascending_iterate_filter: + "ascending (\n. filter_exp (lnot b) (union_up_to_n (if_then b C) S n))" + by (metis ascendingI iterate_sem_assume_increasing iterate_sem_assume_increasing_union_up_to) + + +theorem while_general: + assumes "\n. \ {P n} if_then b C {P (Suc n)}" + and "\n. \ {P n} Assume (lnot b) {Q n}" + and "upwards_closed Q Q_inf" + shows "\ {P 0} while_cond b C {conj Q_inf (holds_forall (lnot b))}" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "P 0 S" + then have "\n. P n (iterate_sem n (if_then b C) S)" + by (meson assms(1) indexed_invariant_then_power) + then have "\n. Q n (filter_exp (lnot b) (union_up_to_n (if_then b C) S n))" + by (metis assms(2) assume_sem filter_exp_def hyper_hoare_triple_def iterate_sem_assume_increasing_union_up_to) + moreover have "ascending (\n. filter_exp (lnot b) (union_up_to_n (if_then b C) S n))" + by (simp add: ascending_iterate_filter) + ultimately have "Q_inf (sem (while_cond b C) S)" + by (metis (no_types, lifting) SUP_cong assms(3) filter_exp_union_general iterate_sem_assume_increasing_union_up_to sem_while_with_if upwards_closed_def) + then show "Logic.conj Q_inf (holds_forall (lnot b)) (sem (while_cond b C) S)" + by (simp add: conj_def filter_exp_def holds_forall_def sem_while_with_if) +qed + +definition while_loop_assertion_n where + "while_loop_assertion_n C S0 n S \ (S = union_up_to_n C S0 n)" + +definition while_loop_assertion_inf where + "while_loop_assertion_inf C S0 S \ (S = (\n. union_up_to_n C S0 n))" + +(* Probably could have completeness with this? *) +lemma while_loop_assertion_upwards_closed: + "upwards_closed (while_loop_assertion_n C S0) (while_loop_assertion_inf C S0)" +proof (rule upwards_closedI) + fix S assume asm0: "ascending S" "\n. while_loop_assertion_n C S0 n (S n)" + then have "\n. S n = union_up_to_n C S0 n" + by (simp add: while_loop_assertion_n_def) + then have "\ (range S) = (\n. union_up_to_n C S0 n)" + by auto + then show "while_loop_assertion_inf C S0 (\ (range S))" + by (simp add: while_loop_assertion_inf_def) +qed + +(* Each element is either always in the sets, or never in the sets, from some point *) +definition converges_sets where + "converges_sets S \ (\x. \n. (\m. m \ n \ (x \ S m)) \ (\m. m \ n \ (x \ S m)))" + +lemma converges_setsI: + assumes "\x. \n. (\m. m \ n \ (x \ S m)) \ (\m. m \ n \ (x \ S m))" + shows "converges_sets S" + by (simp add: assms converges_sets_def) + +lemma ascending_converges: + assumes "ascending S" + shows "converges_sets S" +proof (rule converges_setsI) + fix x + show "\n. (\m\n. x \ S m) \ (\m\n. x \ S m)" + proof (cases "x \ (\n. S n)") + case True + then show ?thesis + by (meson ascending_def assms in_mono) + qed (blast) +qed + +(* Set becomes smaller *) +definition descending :: "(nat \ 'b set) \ bool" where + "descending S \ (\n m. n \ m \ S n \ S m)" + +lemma descending_converges: + assumes "descending S" + shows "converges_sets S" +proof (rule converges_setsI) + fix x + show "\n. (\m\n. x \ S m) \ (\m\n. x \ S m)" + proof (cases "x \ (\n. S n)") + case False + then show ?thesis + by (meson assms descending_def in_mono) + qed (blast) +qed + + +definition limit_sets where + "limit_sets S = {x |x. \n. \m. m \ n \ (x \ S m)}" + +lemma in_limit_sets: + "x \ limit_sets S \ (\n. \m. m \ n \ (x \ S m))" + by (simp add: limit_sets_def) + +lemma ascending_limits_union: + assumes "ascending S" + shows "limit_sets S = (\n. S n)" (is "?A = ?B") +proof + show "?A \ ?B" using limit_sets_def[of S] by auto + show "?B \ ?A" + proof + fix x assume "x \ ?B" + then obtain n where "x \ S n" + by blast + then have "\m. m \ n \ (x \ S m)" + by (meson ascending_def assms subsetD) + then show "x \ ?A" + using limit_sets_def[of S] by auto + qed +qed + +lemma descending_limits_union: + assumes "descending S" + shows "limit_sets S = (\n. S n)" (is "?A = ?B") +proof + show "?B \ ?A" using limit_sets_def[of S] by fastforce + show "?A \ ?B" + proof + fix x assume "x \ ?A" + then obtain n where "\m. m \ n \ (x \ S m)" + using limit_sets_def[of S] by blast + then have "\m. m < n \ (x \ S m)" + by (meson assms descending_def lessI less_imp_le_nat subsetD) + then show "x \ ?B" + by (meson INT_I \\m\n. x \ S m\ linorder_not_le) + qed +qed + + + +definition t_closed where + "t_closed P P_inf \ (\S. converges_sets S \ (\n. P n (S n)) \ P_inf (limit_sets S))" + +lemma t_closed_implies_u_closed: + assumes "t_closed P P_inf" + shows "upwards_closed P P_inf" +proof (rule upwards_closedI) + fix S assume "ascending S" "\n. P n (S n)" + then have "converges_sets S" + using ascending_converges by blast + then show "P_inf (\ (range S))" + by (metis \\n. P n (S n)\ \ascending S\ ascending_limits_union assms t_closed_def) +qed + +(* forall assertions *) +definition downwards_closed where + "downwards_closed P_inf \ (\S S'. S \ S' \ P_inf S' \ P_inf S)" + +(* Slight change compared to Ellora paper *) +definition d_closed where + "d_closed P P_inf \ t_closed P P_inf \ downwards_closed P_inf" + +lemma converges_to_merged: + assumes "\x. x \ S_inf \ \n. \m. m \ n \ (x \ S (m::nat))" + and "\x. x \ S_inf \ \n. \m. m \ n \ (x \ S m)" + shows "converges_sets S \ limit_sets S = S_inf" +proof (rule conjI) + show "converges_sets S" using converges_setsI assms by metis + show "limit_sets S = S_inf" (is "?A = ?B") + proof + show "?B \ ?A" + by (simp add: assms(1) limit_sets_def subsetI) + show "?A \ ?B" + proof + fix x assume "x \ ?A" + then obtain n where n_def: "\m. m \ n \ (x \ S m)" + using in_limit_sets by metis + show "x \ ?B" + proof (rule ccontr) + assume "x \ S_inf" + then obtain n' where "\m. m \ n' \ (x \ S m)" + using assms(2) by presburger + then have "x \ S (max n n') \ x \ S (max n n')" + using n_def by fastforce + then show False by blast + qed + qed + qed +qed + +lemma ascending_union_up: + "ascending (\n. union_up_to_n C S n)" + by (simp add: ascending_def union_up_to_increasing) + +(* actually ascending... *) +lemma converges_union: + "converges_sets (\n. union_up_to_n C S n) \ limit_sets (\n. union_up_to_n C S n) = (\n. union_up_to_n C S n)" +proof (rule converges_to_merged) + fix x + show "x \ \ (range (union_up_to_n C S)) \ \n. \m\n. x \ union_up_to_n C S m" + by (meson UN_iff subset_eq union_up_to_increasing) + show "x \ \ (range (union_up_to_n C S)) \ \n. \m\n. x \ union_up_to_n C S m" + by blast +qed + + +theorem while_d: + assumes "\n. \ {P n} if_then b C {P (Suc n)}" + and "upwards_closed P P_inf" + and "\n. downwards_closed (P n)" \\Satisfied by hyper-assertions that do not existentially quantify over states\ + shows "\ {P 0} while_cond b C {conj P_inf (holds_forall (lnot b))}" + using assms(1) +proof (rule while_general) + show "upwards_closed P P_inf" + using assms(2) by blast + fix n show "\ {P n} Assume (lnot b) {P n}" + proof (rule hyper_hoare_tripleI) + fix S assume "P n S" + moreover have "sem (Assume (lnot b)) S \ S" + by (metis assume_sem member_filter subsetI) + ultimately show "P n (sem (Assume (lnot b)) S)" + by (meson assms(3) downwards_closed_def) + qed +qed + + + +lemma in_union_up_to: + "x \ union_up_to_n C S n \ (\m. m \ n \ x \ iterate_sem m C S)" +proof (induct n) + case (Suc n) + then show ?case + by (metis UnCI UnE le_SucE le_SucI order_refl union_up_to_n.simps(2)) +qed (simp) + + +theorem rule_while_terminates_strong: + assumes "\n. n < m \ \ {P n} if_then b C {P (Suc n)}" + and "\S. P m S \ holds_forall (lnot b) S" + shows "\ {P 0} while_cond b C {P m}" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "P 0 S" + let ?S = "iterate_sem m (if_then b C) S" + let ?S' = "iterate_sem (Suc m) (if_then b C) S" + have "P m ?S" + using asm0 assms(1) indexed_invariant_then_power_bounded by blast + then have "holds_forall (lnot b) ?S" + using assms(2) by auto + moreover have "sem (while_cond b C) S = filter_exp (lnot b) (\n. iterate_sem n (Assume b ;; C) S)" + by (simp add: assume_sem filter_exp_def sem_seq sem_while while_cond_def) + + +(* this is constant *) + then have "P m (filter_exp (lnot b) (union_up_to_n (Assume b;; C) S m) \ iterate_sem m (Assume b;; C) S)" + by (metis \P m (iterate_sem m (if_then b C) S)\ iterate_sem_equiv) + + moreover have "iterate_sem m (Assume b;; C) S \ filter_exp (lnot b) (union_up_to_n (Assume b;; C) S m)" + proof + fix x assume "x \ iterate_sem m (Assume b;; C) S" + then have "x \ union_up_to_n (Assume b;; C) S m" + by (metis UnCI union_up_to_n.elims) + then have "x \ ?S" + by (simp add: \x \ iterate_sem m (Assume b ;; C) S\ iterate_sem_equiv) + then have "lnot b (snd x)" + by (metis calculation(1) holds_forall_def) + then show "x \ filter_exp (lnot b) (union_up_to_n (Assume b;; C) S m)" + using \x \ union_up_to_n (Assume b;; C) S m\ filter_exp_def + by force + qed + moreover have "filter_exp (lnot b) (\n. iterate_sem n (Assume b ;; C) S) + = filter_exp (lnot b) (union_up_to_n (Assume b;; C) S m)" + proof - + have "\n. n > m \ iterate_sem n (Assume b ;; C) S = {}" + proof - + fix n show "n > m \ iterate_sem n (Assume b ;; C) S = {}" + proof (induct "n - m - 1") + case 0 + then show ?case + by (metis (no_types, lifting) UnCI calculation(1) false_then_empty_later holds_forall_def iterate_sem_equiv) + next + case (Suc x) + then show ?case + by (metis (no_types, lifting) UnCI calculation(1) false_then_empty_later holds_forall_def iterate_sem_equiv) + qed + qed + moreover have "union_up_to_n (Assume b;; C) S m = (\n. union_up_to_n (Assume b;; C) S n)" (is "?A = ?B") + proof + show "?B \ ?A" + proof + fix x assume "x \ ?B" + then obtain n where "x \ union_up_to_n (Assume b;; C) S n" + by blast + then show "x \ ?A" + by (metis calculation empty_iff in_union_up_to linorder_not_le) + qed + qed (blast) + then have "(\n. iterate_sem n (Assume b ;; C) S) = union_up_to_n (Assume b;; C) S m" + by (simp add: union_union_up_to_n_equiv) + then show ?thesis + by auto + qed + ultimately show "P m (sem (while_cond b C) S)" + by (simp add: \sem (while_cond b C) S = filter_exp (lnot b) (\n. iterate_sem n (Assume b ;; C) S)\ sup.absorb1) +qed + + +lemma false_state_in_if_then: + assumes "\ \ S" + and "\ b (snd \)" + shows "\ \ sem (if_then b C) S" +proof - + have "\ \ sem (Assume (lnot b)) S" + by (metis SemAssume assms(1) assms(2) in_sem lnot_def prod.collapse) + then show ?thesis + by (simp add: assume_sem filter_exp_def if_then_sem) +qed + +lemma false_state_in_while_cond_aux: + assumes "\ \ S" + and "\ b (snd \)" + shows "\ \ iterate_sem n (if_then b C) S" +proof (induct n) + case 0 + then show ?case + by (simp add: assms(1)) +next + case (Suc n) + then show ?case + by (simp add: assms(2) false_state_in_if_then) +qed + +lemma false_state_in_while_cond: + assumes "\ \ S" + and "\ b (snd \)" + shows "\ \ sem (while_cond b C) S" +proof - + have "\ \ (\n. iterate_sem n (if_then b C) S)" + by (simp add: assms(1) assms(2) false_state_in_while_cond_aux) + then show ?thesis using sem_while_with_if[of b C S] filter_exp_def lnot_def + by (metis assms(2) comp_apply member_filter) +qed + +theorem while_exists: + assumes "\\. \ { P \ } while_cond b C { Q \ }" + shows "\ { (\S. \\ \ S. \ b (snd \) \ P \ S) } while_cond b C { (\S. \\ \ S. Q \ S) }" +proof (rule hyper_hoare_tripleI) + fix S assume "\\\S. \ b (snd \) \ P \ S" + then obtain \ where asm0: "\\S" "\ b (snd \) \ P \ S" by blast + then have "Q \ (sem (while_cond b C) S)" + using assms hyper_hoare_tripleE by blast + then show "\\\sem (while_cond b C) S. Q \ (sem (while_cond b C) S)" + using asm0(1) asm0(2) false_state_in_while_cond by blast +qed + +lemma sem_while_cond_union_up_to: + "sem (while_cond b C) S = filter_exp (lnot b) (\n. union_up_to_n (if_then b C) S n)" + by (simp add: sem_while_with_if union_union_up_to_n_equiv) + +lemma iterate_sem_sum: + "iterate_sem n C (iterate_sem m C S) = iterate_sem (n + m) C S" + by (induct n) simp_all + + +lemma unroll_while_sem: + "sem (while_cond b C) (iterate_sem n (if_then b C) S) = sem (while_cond b C) S" +proof - + let ?S = "iterate_sem n (if_then b C) S" + have "filter_exp (lnot b) (\m. iterate_sem m (if_then b C) S) = filter_exp (lnot b) (\m. iterate_sem (n + m) (if_then b C) S)" (is "?A = ?B") + proof + show "?A \ ?B" + proof + fix x assume "x \ ?A" + then obtain m where "x \ iterate_sem m (if_then b C) S" "\ b (snd x)" + using filter_exp_def lnot_def + by (metis (no_types, lifting) UN_iff comp_apply member_filter) + then have "x \ iterate_sem (n + m) (if_then b C) S" + using false_state_in_while_cond_aux[of x "iterate_sem m (if_then b C) S" b n C] iterate_sem_sum[of n "if_then b C" m S] + by blast + then have "x \ (\m. iterate_sem (n + m) (if_then b C) S)" + by blast + then show "x \ ?B" + by (metis \x \ filter_exp (lnot b) (\m. iterate_sem m (if_then b C) S)\ filter_exp_def member_filter) + qed + show "?B \ ?A" + proof + fix x assume "x \ ?B" + then obtain m where "x \ iterate_sem (n + m) (if_then b C) S" "\ b (snd x)" + by (metis (no_types, lifting) UN_iff comp_apply filter_exp_def lnot_def member_filter) + then show "x \ ?A" + by (metis (no_types, lifting) UNIV_I UN_iff \x \ filter_exp (lnot b) (\m. iterate_sem (n + m) (if_then b C) S)\ filter_exp_def member_filter) + qed + qed + then show ?thesis + using iterate_sem_sum[of _ "if_then b C" n S] sem_while_with_if[of b C S] sem_while_with_if[of b C ?S] + by (simp add: add.commute) +qed + + +theorem while_unroll: + assumes "\n. n < m \ \ {P n} if_then b C {P (Suc n)}" + and "\ {P m} while_cond b C {Q}" + shows "\ {P 0} while_cond b C {Q}" +proof (rule hyper_hoare_tripleI) + fix S assume "P 0 S" + let ?S = "iterate_sem m (if_then b C) S" + have "(\n. n < m \ (\ {P n} if_then b C {P (Suc n)})) \ P m ?S" + proof (induct m) + case 0 + then show ?case + by (simp add: \P 0 S\) + next + case (Suc m) + then show ?case + by (simp add: hyper_hoare_triple_def) + qed + then have "P m ?S" using assms(1) + by blast + then have "Q (sem (while_cond b C) ?S)" + using assms(2) hyper_hoare_tripleE by blast + then show "Q (sem (while_cond b C) S)" + by (metis unroll_while_sem) +qed + + + + + + + + +text \Deriving LoopExit from NormalWhile, and ForLoop from LoopExit and Unroll\ + +lemma while_desugared_easy: + assumes "\n. \ {I n} Assume b;; C {I (Suc n)}" + and "\ { natural_partition I } Assume (lnot b) { Q }" + shows "\ {I 0} while_cond b C { Q }" + by (metis assms(1) assms(2) seq_rule while_cond_def while_rule) + + +corollary loop_exit: + assumes "entails P (holds_forall (lnot b))" + shows "\ {P} while_cond b C {P}" +proof - + have "\ {(if (0::nat) = 0 then P else emp)} while_cond b C {P}" + proof (rule while_desugared_easy[of "\(n::nat). if n = 0 then P else emp" b C P]) + show "\ {natural_partition (\(n::nat). if n = 0 then P else emp)} Assume (lnot b) {P}" + proof (rule hyper_hoare_tripleI) + fix S assume asm0: "natural_partition (\(n::nat). if n = 0 then P else emp) S" + then obtain F where "S = (\(n::nat). F n)" "\(n::nat). (\(n::nat). if n = 0 then P else emp) n (F n)" + using natural_partitionE by blast + then have "\n. F (Suc n) = {}" + by (metis (mono_tags, lifting) emp_def old.nat.distinct(2)) + moreover have "S = F 0" + proof + show "S \ F 0" + proof + fix x assume "x \ S" + then obtain n where "x \ F n" + using \S = \ (range F)\ by blast + then show "x \ F 0" + by (metis calculation empty_iff gr0_implies_Suc zero_less_iff_neq_zero) + qed + show "F 0 \ S" + using \S = \ (range F)\ by blast + qed + ultimately have "P S" + using \\n. (if n = 0 then P else emp) (F n)\ by presburger + then show "P (sem (Assume (lnot b)) S)" + by (metis assms entailsE sem_assume_low_exp(1)) + qed + fix n :: nat + show "\ {(if n = 0 then P else emp)} Assume b ;; C {if Suc n = 0 then P else emp}" + proof (rule hyper_hoare_tripleI) + fix S assume asm0: "(if n = 0 then P else emp) S" + then show "(if Suc n = 0 then P else emp) (sem (Assume b ;; C) S)" + by (metis (mono_tags, lifting) assms emp_def entailsE holds_forall_empty lnot_involution nat.distinct(1) sem_assume_low_exp_seq(2)) + qed + qed + then show ?thesis + by fastforce +qed + +corollary for_loop: + assumes "\n. n < m \ \ {P n} if_then b C {P (Suc n)}" + and "entails (P m) (holds_forall (lnot b))" + shows "\ {P 0} while_cond b C {P m}" + using assms(1) +proof (rule while_unroll) + show "\ {P m} while_cond b C {P m}" + using assms(2) loop_exit by blast +qed + + +end \ No newline at end of file diff --git a/thys/HyperHoareLogic/PaperResults.thy b/thys/HyperHoareLogic/PaperResults.thy new file mode 100644 --- /dev/null +++ b/thys/HyperHoareLogic/PaperResults.thy @@ -0,0 +1,540 @@ +theory PaperResults + imports Loops SyntacticAssertions Compositionality TotalLogic ExamplesCompositionality +begin + +section \Summary of the Results from the Paper\ + +text \This file contains the formal results mentioned the paper. It is organized in the same order +and with the same structure as the paper. +\<^item> You can use the panel "Sidekick" on the right to see and navigate the structure of the file, via sections and subsections. +\<^item> You can ctrl+click on terms to jump to their definition. +\<^item> After jumping to another location, you can come back to the previous location by clicking the + green left arrow, on the right side of the menu above.\ + + + +subsection \3: Hyper Hoare Logic\ + +subsubsection \3.1: Language and Semantics\ + +text \The programming language is defined in the file Language.thy: +\<^item> The type of program state (definition 1) is \<^typ>\('pvar, 'pval) pstate\ + (<-- you can ctrl+click on the name \pstate\ above to jump to its definition). +\<^item> Program commands (definition 1) are defined via the type \<^typ>\('var, 'val) stmt\. +\<^item> The big-step semantics (figure 9) is defined as \<^const>\single_sem\. We also use the notation \<^term>\\C, \\ \ \'\.\ + +subsubsection \3.2: Hyper-Triples, Formally\ + +text \ +\<^item> Extended states (definition 2) are defined as \<^typ>\('lvar, 'lval, 'pvar, 'pval) state\ (file Language.thy). +\<^item> Hyper-assertions (definition 3) are defined as \<^typ>\('lvar, 'lval, 'pvar, 'pval) state hyperassertion\ (file Logic.thy). +\<^item> The extended semantics (definition 4) is defined as \<^const>\sem\ (file Language.thy). +\<^item> Lemma 1 is shown and proven below. +\<^item> Hyper-triples (definition 5) are defined as \<^const>\hyper_hoare_triple\ (file Logic.thy). + We also use the notation \<^term>\\ {P} C {Q}\.\ + +lemma lemma1: + "sem C (S1 \ S2) = sem C S1 \ sem C S2" + "S \ S' \ sem C S \ sem C S'" + "sem C (\x. f x) = (\x. sem C (f x))" + "sem Skip S = S" + "sem (C1;; C2) S = sem C2 (sem C1 S)" + "sem (If C1 C2) S = sem C1 S \ sem C2 S" + "sem (While C) S = (\n. iterate_sem n C S)" + using sem_if sem_seq sem_union sem_skip sem_union_general sem_monotonic sem_while by metis+ + +subsubsection \3.3: Core Rules\ + +text \The core rules (from figure 2) are defined in the file Logic.thy as \<^const>\syntactic_HHT\. +We also use the notation \<^term>\\ {P} C {Q}\. Operators \\\ (definition 6) and \\\ (definition 7) +are defined as \<^const>\join\ and \<^const>\natural_partition\, respectively.\ + +subsubsection \3.4: Soundness and Completeness\ + +text \Theorem 1: Soundness\ +theorem thm1_soundness: + assumes "\ {P} C {Q}" + shows "\ {P} C {Q}" + using assms soundness by auto + +text \Theorem 2: Completeness\ +theorem thm2_completeness: + assumes "\ {P} C {Q}" + shows "\ {P} C {Q}" + using assms completeness by auto + +subsubsection \3.5: Expressivity of Hyper-Triples\ + +text \Program hyperproperties (definition 8) are defined in the file ProgramHyperproperties as +the type \<^typ>\('pvar, 'pval) program_hyperproperty\, which is syntactic sugar for the type +\<^typ>\(('pvar, 'pval) pstate \ ('pvar, 'pval) pstate) set \ bool\. As written in the paper (after the definition), +this type is equivalent to the type \<^typ>\((('pvar, 'pval) pstate \ ('pvar, 'pval) pstate) set) set\. +The satisfiability of program hyperproperties is defined via the function \<^const>\hypersat\.\ + +text \Theorem 3: Expressing hyperproperties as hyper-triples\ +theorem thm3_expressing_hyperproperties_as_hyper_triples: + fixes to_lvar :: "'pvar \ 'lvar" + fixes to_lval :: "'pval \ 'lval" + fixes H :: "('pvar, 'pval) program_hyperproperty" + assumes "injective to_lvar" \\The cardinality of \<^typ>\'lvar\ is at least the cardinality of \<^typ>\'pvar\.\ + and "injective to_lval" \\The cardinality of \<^typ>\'lval\ is at least the cardinality of \<^typ>\'pval\.\ + shows "\P Q::('lvar, 'lval, 'pvar, 'pval) state hyperassertion. (\C. hypersat C H \ \ {P} C {Q})" + using assms proving_hyperproperties + by blast + +text \Theorem 4: Expressing hyper-triples as hyperproperties\ +theorem thm4_expressing_hyper_triples_as_hyperproperties: + "\ {P} C {Q} \ hypersat C (hyperprop_hht P Q)" + by (simp add: any_hht_hyperprop) + +text \Theorem 5: Disproving hyper-triples\ +theorem thm5_disproving_triples: + "\ \ {P} C {Q} \ (\P'. sat P' \ entails P' P \ \ {P'} C {\S. \ Q S})" + using disproving_triple by auto + + + + +subsection \4: Syntactic Rules\ + +subsubsection \4.1: Syntactic Hyper-Assertions\ + +text \Syntactic hyper-expressions and hyper-assertions (definition 9) are defined in the file +SyntacticAssertions.thy as \<^typ>\'val exp\ and \<^typ>\'val assertion\ respectively, where 'val is the +type of both logical and program values. Note that we use de Bruijn indices (i.e, natural numbers) +for states and variables bound by quantifiers.\ + +subsubsection \4.2: Syntactic Rules for Deterministic and Non-Deterministic Assignments.\ + +text \We prove semantic versions of the syntactic rules from subsection 4 (figure 3). +We use \<^const>\interp_assert\ to convert a syntactic hyper-assertion into a semantic one, because +our hyper-triples require semantic hyper-assertions. Similarly, we use \<^const>\interp_pexp\ to convert +a syntactic program expression into a semantic one. +\<^term>\transform_assign x e P\ and \<^term>\transform_havoc x P\ correspond to \A\<^sup>e\<^sub>x\ and \H\<^sub>x\ from definition 10.\ + +text \Rule AssignS from figure 3\ +proposition AssignS: + "\ { interp_assert (transform_assign x e P) } Assign x (interp_pexp e) {interp_assert P}" + using completeness rule_assign_syntactic by blast + +text \Rule HavocS from figure 3\ +proposition HavocS: + "\ { interp_assert (transform_havoc x P) } Havoc x {interp_assert P}" + using completeness rule_havoc_syntactic by blast + +subsubsection \4.3: Syntactic Rules for Assume Statements\ + +text \\<^const>\transform_assume\ corresponds to \\\<^sub>b\ (definition 11).\ + +text \Rule AssumeS from figure 3\ +proposition AssumeS: + "\ { interp_assert (transform_assume (pbexp_to_assertion 0 b) P) } Assume (interp_pbexp b) {interp_assert P}" + using completeness rule_assume_syntactic by blast + +text \As before, we use \<^const>\interp_pbexp\ to convert the syntactic program Boolean expression b into +a semantic one. Similarly, \<^term>\pbexp_to_assertion 0 b\ converts the syntactic program Boolean expression +p into a syntactic hyper-assertion. The number 0 is a de Bruijn index, which corresponds to the closest +quantified state. For example, the hyper-assertion \\<\>. \(a)=\(b) \ (\<\'>. \(x) \ \'(y))\ would +be written as \\. 0(a)=0(b) \ (\. 1(x) \ 0(y))\ with de Bruijn indices. Thus, one can think of +\pbexp_to_assertion 0 b\ as \b(\)\, where \\\ is simply the innermost quantified state.\ + + + + +subsection \5: Proof Principles for Loops\ + +text \We show in the following our proof rules for loops, presented in figure 5.\ + +text \Rule WhileDesugared from figure 5\ +theorem while_desugared: + assumes "\n. \ {I n} Assume b;; C {I (Suc n)}" + and "\ { natural_partition I } Assume (lnot b) { Q }" + shows "\ {I 0} while_cond b C { Q }" + by (metis completeness soundness assms(1) assms(2) seq_rule while_cond_def while_rule) + +text \This result uses the following constructs: +\<^item> \<^term>\natural_partition I\ corresponds to the \\\ operator from definition 7. +\<^item> \<^term>\lnot b\ negates b. +\<^item> \<^term>\while_cond b C\ is defined as \<^term>\While (Assume b;; C);; Assume (lnot b)\.\ + +text \Rule WhileSync from figure 5 (presented in subsubsection 5.1)\ +lemma WhileSync: + assumes "entails I (low_exp b)" + and "\ {conj I (holds_forall b)} C {I}" + shows "\ {conj I (low_exp b)} while_cond b C {conj (disj I emp) (holds_forall (lnot b))}" + using WhileSync_simpler entail_conj completeness soundness assms(1) assms(2) entails_refl + consequence_rule[of "conj I (holds_forall b)" "conj I (holds_forall b)" I "conj I (low_exp b)"] by blast + +text \This result uses the following constructs: +\<^item> \<^term>\conj A B\ corresponds to the hyper-assertion \A \ B\. +\<^item> \<^term>\holds_forall b\ corresponds to \box(b)\. +\<^item> \<^term>\low_exp b\ corresponds to \low(b)\. +\<^item> \<^term>\disj A B\ corresponds to the hyper-assertion \A \ B\. +\<^item> \<^term>\emp\ checks whether the set of states is empty.\ + +text \Rule IfSync from figure 5 (presented in subsubsection 5.1)\ +theorem IfSync: + assumes "entails P (low_exp b)" + and "\ {conj P (holds_forall b)} C1 {Q}" + and "\ {conj P (holds_forall (lnot b))} C2 {Q}" + shows "\ {P} if_then_else b C1 C2 {Q}" + using completeness soundness consequence_rule[of _ "conj P (low_exp b)" Q] assms(1) entail_conj entails_refl assms if_synchronized by metis + +text \This result uses the following construct: +\<^item> \<^term>\if_then_else b C1 C2\ is syntactic sugar for \<^term>\If (Assume b;; C1) (Assume (lnot b);; C2)\.\ + +text \Rule \While-\*\*\ from figure 5 (presented in subsubsection 5.2)\ +theorem while_forall_exists: + assumes "\ {I} if_then b C {I}" + and "\ {I} Assume (lnot b) {interp_assert Q}" + and "no_forall_state_after_existential Q" + shows "\ {I} while_cond b C {interp_assert Q}" + using consequence_rule[of I I "conj (interp_assert Q) (holds_forall (lnot b))" "interp_assert Q"] + using completeness soundness while_forall_exists_simpler assms entail_conj_weaken entails_refl by metis + +text \This result uses the following constructs: +\<^item> \<^term>\if_then b C\ is syntactic sugar for \<^term>\If (Assume b;; C) (Assume (lnot b))\. +\<^item> \<^term>\no_forall_state_after_existential Q\ holds iff there is no universal state quantifier \\\_\\ after any \\\ in Q.\ + +text \Rule \While-\\ from figure 5 (presented in subsubsection 5.3)\ +theorem while_loop_exists: + assumes "\v. \ { (\S. \\\S. e (snd \) = v \ b (snd \) \ P \ S) } if_then b C { (\S. \\\S. lt (e (snd \)) v \ P \ S) }" + and "\\. \ { P \ } while_cond b C { Q \ }" + and "wfP lt" + shows "\ { (\S. \\\S. P \ S) } while_cond b C { (\S. \\\S. Q \ S)}" + using completeness soundness exists_terminates_loop assms by blast + +text \\<^term>\wfP lt\ in this result ensures that the binary operator \lt\ is well-founded. +\e\ is a function of a program state, which must decrease after each iteration.\ + + + + +subsection \Appendix A: Technical Definitions Omitted from the Paper\ + +text \The big-step semantics (figure 9) is defined as \<^const>\single_sem\. We also use the notation \<^term>\\C, \\ \ \'\. +The following definitions are formalized in the file SyntacticAssertions.thy: +\<^item> Evaluation of hyper-expressions (definition 12): \<^const>\interp_exp\. +\<^item> Satisfiability of hyper-assertions (definition 12): \<^const>\sat_assertion\. +\<^item> Syntactic transformation for deterministic assignments (definition 13): \<^const>\transform_assign\. +\<^item> Syntactic transformation for non-deterministic assignments (definition 14): \<^const>\transform_havoc\. +\<^item> Syntactic transformation for assume statements. (definition 15): \<^const>\transform_assume\.\ + + + +subsection \Appendix C: Expressing Judgments of Hoare Logics as Hyper-Triples\ + +subsubsection \Appendix C.1: Overapproximate Hoare Logics\ + +text \The following judgments are defined in the file Expressivity.thy as follows: +\<^item> Definition 16 (Hoare Logic): \<^term>\HL P C Q\. +\<^item> Definition 17 (Cartesian Hoare Logic): \<^term>\CHL P C Q\.\ + +text \Proposition 1: HL triples express hyperproperties\ +proposition prop_1_HL_expresses_hyperproperties: + "\H. (\C. hypersat C H \ HL P C Q)" + using HL_expresses_hyperproperties by blast + +text \Proposition 2: Expressing HL in Hyper Hoare Logic\ +proposition prop_2_expressing_HL_in_HHL: + "HL P C Q \ (hyper_hoare_triple (over_approx P) C (over_approx Q))" + using encoding_HL by auto + +text \Proposition 3: CHL triples express hyperproperties\ +proposition prop_3_CHL_is_hyperproperty: + "hypersat C (CHL_hyperprop P Q) \ CHL P C Q" + using CHL_hyperproperty by fast + +text \Proposition 4: Expressing CHL in Hyper Hoare Logic\ +proposition prop_4_encoding_CHL_in_HHL: + assumes "not_free_var_of P x" + and "not_free_var_of Q x" + and "injective from_nat" + shows "CHL P C Q \ \ {encode_CHL from_nat x P} C {encode_CHL from_nat x Q}" + using encoding_CHL assms by fast + +text \The function \from_nat\ gives us a way to encode numbers from 1 to k as logical values. +Moreover, note that we represent k-tuples implicitly, as mappings of type \<^typ>\'a \ 'b\: When the type +\'a\ has k elements, a function of type \<^typ>\'a \ 'b\ corresponds to a k-tuple of elements of type 'b. +This representation is more convenient to work with, and more general, since it also captures infinite sequences.\ + +subsubsection \Appendix C.2: Underapproximate Hoare Logics\ + +text \The following judgments are defined in the file Expressivity.thy as follows: +\<^item> Definition 18 (Incorrectness Logic): \<^term>\IL P C Q\. +\<^item> Definition 19 (k-Incorrectness Logic): \<^term>\RIL P C Q\. +\<^item> Definition 20 (Forward Underapproximation): \<^term>\FU P C Q\. +\<^item> Definition 21 (k-Forward Underapproximation): \<^term>\RFU P C Q\.\ + +text \RIL is the old name of k-IL, and RFU is the old name of k-FU.\ + +text \Proposition 5: IL triples express hyperproperties\ +proposition prop_5_IL_hyperproperties: + "IL P C Q \ IL_hyperprop P Q (set_of_traces C)" + using IL_expresses_hyperproperties by fast + +text \Proposition 6: Expressing IL in Hyper Hoare Logic\ +proposition prop_6_expressing_IL_in_HHL: + "IL P C Q \ (hyper_hoare_triple (under_approx P) C (under_approx Q))" + using encoding_IL by fast + +text \Proposition 7: k-IL triples express hyperproperties\ +proposition prop_7_kIL_hyperproperties: + "hypersat C (RIL_hyperprop P Q) \ RIL P C Q" + using RIL_expresses_hyperproperties by fast + +text \Proposition 8: Expressing k-IL in Hyper Hoare Logic\ +proposition prop_8_expressing_kIL_in_HHL: + fixes x :: "'lvar" + assumes "\l l' \. (\i. (l i, \ i)) \ P \ (\i. (l' i, \ i)) \ P" (* P does not depend on logical variables *) + and "injective (indexify :: (('a \ ('pvar \ 'pval)) \ 'lval))" (* |lval| \ |PStates^k| *) + and "c \ x" + and "injective from_nat" + and "not_free_var_of (P :: ('a \ ('lvar \ 'lval) \ ('pvar \ 'pval)) set) x \ not_free_var_of P c" + and "not_free_var_of Q x \ not_free_var_of Q c" + shows "RIL P C Q \ \ {pre_insec from_nat x c P} C {post_insec from_nat x c Q}" + using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) by (rule encoding_RIL) + +proposition FU_hyperproperties: + "hypersat C (hyperprop_FU P Q) \ FU P C Q" + by (rule FU_expresses_hyperproperties) + +text \Proposition 9: Expressing FU in Hyper Hoare Logic\ +proposition prop_9_expressing_FU_in_HHL: + "FU P C Q \ \ {encode_FU P} C {encode_FU Q}" + by (rule encoding_FU) + +text \Proposition 10: k-FU triples express hyperproperties\ +proposition prop_10_kFU_expresses_hyperproperties: + "hypersat C (RFU_hyperprop P Q) \ RFU P C Q" + by (rule RFU_captures_hyperproperties) + +text \Proposition 11: Expressing k-FU in Hyper Hoare Logic\ +proposition prop_11_encode_kFU_in_HHL: + assumes "not_free_var_of P x" + and "not_free_var_of Q x" + and "injective from_nat" + shows "RFU P C Q \ \ {encode_RFU from_nat x P} C {encode_RFU from_nat x Q}" + using assms + by (rule encode_RFU) + +subsubsection \Appendix C.3: Beyond Over- and Underapproximation\ + +text \The following judgment is defined in the file Expressivity.thy as follows: +\<^item> Definition 22 (k-Universal Existential): \<^term>\RUE P C Q\. +Note that RUE is the old name of k-UE.\ + +text \Proposition 12: k-UE triples express hyperproperties\ +proposition prop_12_kUE_expresses_hyperproperty: + "RUE P C Q \ hypersat C (hyperprop_RUE P Q)" + by (rule RUE_express_hyperproperties) + +text \Proposition 13: Expressing k-UE in Hyper Hoare Logic\ +proposition prop_13_expressing_kUE_in_HHL: + assumes "injective fn \ injective fn1 \ injective fn2" + and "t \ x" + and "injective (fn :: nat \ 'a)" + and "injective fn1" + and "injective fn2" + and "not_in_free_vars_double {x, t} P" + and "not_in_free_vars_double {x, t} Q" + shows "RUE P C Q \ \ {encode_RUE_1 fn fn1 fn2 x t P} C {encode_RUE_2 fn fn1 fn2 x t Q}" + using assms by (rule encoding_RUE) + + +text \Example 3\ +proposition proving_refinement: + fixes P :: "(('lvar \ 'lval) \ ('pvar \ 'pval)) set \ bool" + and t :: 'pvar + assumes "(one :: 'pval) \ two" \\We assume two distinct program values \one\ and \two\, to represent 1 and 2.\ + and "P = (\S. card S = 1)" + and "Q = (\S. \\\S. snd \ t = two \ (fst \, (snd \)(t := one)) \ S)" + and "not_free_var_stmt t C1" + and "not_free_var_stmt t C2" + shows "refinement C2 C1 \ + \ { P } If (Seq (Assign t (\_. one)) C1) (Seq (Assign t (\_. two)) C2) { Q }" +proof - + have "refinement C2 C1 \ \ { P } If (Seq (Assign t (\_. two)) C2) (Seq (Assign t (\_. one)) C1) { Q }" + using encoding_refinement[of two one P Q t C2 C1] assms by blast + then show ?thesis using rewrite_if_commute by blast +qed + + + + +subsection \Appendix D: Compositionality\ + +subsubsection \Appendix D.1: Compositionality Rules\ + +text \In the following, we show the rules from figure 11, in the order in which they appear.\ + +proposition rule_Linking: + assumes "\\1 (\2 :: ('a, 'b, 'c, 'd) state). fst \1 = fst \2 \ ( \ { (in_set \1 :: (('a, 'b, 'c, 'd) state) hyperassertion) } C { in_set \2 } ) + \ ( \ { (P \1 :: (('a, 'b, 'c, 'd) state) hyperassertion) } C { Q \2 } )" + shows "\ { ((\S. \\1 \ S. P \1 S) :: (('a, 'b, 'c, 'd) state) hyperassertion) } C { (\S. \\2 \ S. Q \2 S) }" + using assms soundness completeness rule_linking by blast + +proposition rule_And: + assumes "\ {P} C {Q}" + and "\ {P'} C {Q'}" + shows "\ {conj P P'} C {conj Q Q'}" + using assms soundness completeness rule_And by metis + +proposition rule_Or: + assumes "\ {P} C {Q}" + and "\ {P'} C {Q'}" + shows "\ {disj P P'} C {disj Q Q'}" + using assms soundness completeness rule_Or by metis + +proposition rule_FrameSafe: + assumes "wr C \ fv F = {}" + and "wf_assertion F" + and "no_exists_state F" + shows "\ {interp_assert F} C {interp_assert F}" + using safe_frame_rule_syntactic assms completeness by metis + +proposition rule_Forall: + assumes "\x. \ {P x} C {Q x}" + shows "\ {forall P} C {forall Q}" + using assms soundness completeness rule_Forall by metis + +proposition rule_IndexedUnion: + assumes "\x. \ {P x} C {Q x}" + shows "\ {general_join P} C {general_join Q}" + using assms soundness completeness rule_IndexedUnion by metis + +proposition rule_Union: + assumes "\ {P} C {Q}" + and "\ {P'} C {Q'}" + shows "\ {join P P'} C {join Q Q'}" + using assms soundness completeness rule_Union by metis + +proposition rule_BigUnion: + fixes P :: "((('a \ 'b) \ ('c \ 'd)) set \ bool)" + assumes "\ {P} C {Q}" + shows "\ {general_union P} C {general_union Q}" + using assms soundness completeness rule_BigUnion by blast + +proposition rule_Specialize: + assumes "\ {interp_assert P} C {interp_assert Q}" + and "indep_of_set b" + and "wf_assertion_aux 0 1 b" + and "wr C \ fv b = {}" + shows "\ { interp_assert (transform_assume b P) } C { interp_assert (transform_assume b Q) }" + using filter_rule_syntactic assms soundness completeness by blast + +text \In the following, \<^term>\entails_with_updates vars P P'\ and \<^term>\invariant_on_updates vars Q\ +respectively correspond to the notions of entailments modulo logical variables and invariance with +respect to logical updates, as described in definition 23.\ + +proposition rule_LUpdate: + assumes "\ {P'} C {Q}" + and "entails_with_updates vars P P'" + and "invariant_on_updates vars Q" + shows "\ {P} C {Q}" + using assms soundness completeness rule_LUpdate by blast + +proposition rule_LUpdateSyntactic: + assumes "\ { (\S. P S \ e_recorded_in_t e t S) } C { Q }" + and "not_fv_hyper t P" + and "not_fv_hyper t Q" + shows "\ { P } C { Q }" + using LUpdateS soundness completeness assms by fast + +proposition rule_AtMost: + assumes "\ {P} C {Q}" + shows "\ {has_superset P} C {has_superset Q}" + using assms soundness completeness rule_AtMost by blast + +(* derived from join *) +proposition rule_AtLeast: + assumes "\ {P} C {Q}" + shows "\ {has_subset P} C {has_subset Q}" + using assms soundness completeness rule_AtLeast by blast + +proposition rule_True: + "\ {P} C {\_. True}" + using rule_True completeness by blast + +proposition rule_False: + "\ { (\_. False) } C {Q}" + using rule_False completeness by blast + +proposition rule_Empty: + "\ { (\S. S = {}) } C { (\S. S = {}) }" + using completeness rule_Empty by blast + + +subsubsection \Appendix D.2: Examples\ + +text \Example shown in figure 12. To see the actual proof, ctrl+click on @{thm composing_monotonicity_and_minimum}.\ +proposition fig_12_composing_monotonicity_and_minimum: + fixes P :: "((('a \ 'b) \ ('c \ 'd)) set \ bool)" + fixes i :: 'a + fixes x y :: 'c + fixes leq :: "'d \ 'd \ bool" + fixes one two :: 'b + assumes "\ { P } C1 { has_minimum x leq }" + and "\ { is_monotonic i x one two leq } C2 { is_monotonic i y one two leq }" + and "\ { (is_singleton :: ((('a \ 'b) \ ('c \ 'd)) set \ bool)) } C2 { is_singleton }" + and "one \ two" \\We use distinct logical values \one\ and \two\ to represent 1 and 2.\ + and "\x. leq x x" \\We assume that \leq\ is a partial order, and thus that it satisfies reflexivity.\ + shows "\ { P } C1 ;; C2 { has_minimum y leq }" + using assms soundness completeness composing_monotonicity_and_minimum by metis + +text \Example shown in figure 13. To see the actual proof, ctrl+click on @{thm composing_GNI_with_SNI}.\ +proposition fig_13_composing_GNI_with_SNI: + fixes h :: 'lvar + fixes l :: 'pvar + assumes "\ { (low l :: (('lvar, 'lval, 'pvar, 'pval) state) hyperassertion) } C2 { low l }" + and "\ { (not_empty :: (('lvar, 'lval, 'pvar, 'pval) state) hyperassertion) } C2 { not_empty }" + and "\ { (low l :: (('lvar, 'lval, 'pvar, 'pval) state) hyperassertion) } C1 { lGNI l h }" + shows "\ { (low l :: (('lvar, 'lval, 'pvar, 'pval) state) hyperassertion) } C1;; C2 { lGNI l h }" + using assms soundness completeness composing_GNI_with_SNI by metis + + + + + +subsection \Appendix E: Termination-Based Reasoning\ + +text \Terminating hyper-triples (definition 24) are defined as \<^const>\total_hyper_triple\, and usually +written \<^term>\\TERM {P} C {Q}\.\ + +theorem rule_Frame: + assumes "\TERM {P} C {Q}" + and "wr C \ fv F = {}" + and "wf_assertion F" + shows "\TERM {conj P (interp_assert F)} C {conj Q (interp_assert F)}" + by (simp add: assms(1) assms(2) assms(3) frame_rule_syntactic) + +theorem rule_WhileSyncTerm: + assumes "\TERM {conj I (\S. \\\S. b (snd \) \ fst \ t = e (snd \))} C {conj (conj I (low_exp b)) (e_smaller_than_t e t lt)}" + and "wfP lt" + and "not_fv_hyper t I" + shows "\TERM {conj I (low_exp b)} while_cond b C {conj I (holds_forall (lnot b))}" + by (meson WhileSyncTot assms(1) assms(2) assms(3)) + + + + +subsection \Appendix H: Synchronous Reasoning over Different Branches\ + +text \Proposition 14: Synchronous if rule\ +proposition prop_14_synchronized_if_rule: + assumes "\ {P} C1 {P1}" + and "\ {P} C2 {P2}" + and "\ {combine from_nat x P1 P2} C {combine from_nat x R1 R2}" + and "\ {R1} C1' {Q1}" + and "\ {R2} C2' {Q2}" + and "not_free_var_hyper x P1" + and "not_free_var_hyper x P2" + and "from_nat 1 \ from_nat 2" \\We can represent 1 and 2 as distinct logical values.\ + and "not_free_var_hyper x R1" + and "not_free_var_hyper x R2" + shows "\ {P} If (Seq C1 (Seq C C1')) (Seq C2 (Seq C C2')) {join Q1 Q2}" + using if_sync_rule assms by meson + + + +end \ No newline at end of file diff --git a/thys/HyperHoareLogic/ProgramHyperproperties.thy b/thys/HyperHoareLogic/ProgramHyperproperties.thy --- a/thys/HyperHoareLogic/ProgramHyperproperties.thy +++ b/thys/HyperHoareLogic/ProgramHyperproperties.thy @@ -1,339 +1,341 @@ section \Expressivity of Hyper Hoare Logic\ -text \In this file, we define program hyperproperties (definition 7), and prove theorem 3.\ +text \In this file, we define program hyperproperties (definition 8), and prove theorems 3 and 4.\ subsection \Program Hyperproperties\ theory ProgramHyperproperties imports Logic begin -text \Definition 7\ +text \Definition 8\ type_synonym 'a hyperproperty = "('a \ 'a) set \ bool" +type_synonym ('pvar, 'pval) program_hyperproperty = "('pvar, 'pval) pstate hyperproperty" + definition set_of_traces where "set_of_traces C = { (\, \') |\ \'. \C, \\ \ \' }" -definition hypersat where +definition hypersat :: "('pvar, 'pval) stmt \ ('pvar, 'pval) program_hyperproperty \ bool" where "hypersat C H \ H (set_of_traces C)" definition copy_p_state where "copy_p_state to_pvar to_lval \ x = to_lval (\ (to_pvar x))" definition recover_p_state where "recover_p_state to_pval to_lvar l x = to_pval (l (to_lvar x))" lemma injective_then_exists_inverse: assumes "injective to_lvar" shows "\to_pvar. (\x. to_pvar (to_lvar x) = x)" proof - let ?to_pvar = "\y. SOME x. to_lvar x = y" have "\x. ?to_pvar (to_lvar x) = x" by (metis (mono_tags, lifting) assms injective_def someI) then show ?thesis by force qed lemma single_step_then_in_sem: assumes "single_sem C \ \'" and "(l, \) \ S" shows "(l, \') \ sem C S" using assms(1) assms(2) in_sem by fastforce lemma in_set_of_traces: "(\, \') \ set_of_traces C \ \C, \\ \ \'" by (simp add: set_of_traces_def) lemma in_set_of_traces_then_in_sem: assumes "(\, \') \ set_of_traces C" and "(l, \) \ S" shows "(l, \') \ sem C S" using in_set_of_traces assms single_step_then_in_sem by metis lemma set_of_traces_same: assumes "\x. to_pvar (to_lvar x) = x" and "\x. to_pval (to_lval x) = x" and "S = {(copy_p_state to_pvar to_lval \, \) |\. True}" shows "{(recover_p_state to_pval to_lvar l, \') |l \'. (l, \') \ sem C S} = set_of_traces C" (is "?A = ?B") proof show "?A \ ?B" proof (rule subsetPairI) fix \ \' assume asm0: "(\, \') \ {(recover_p_state to_pval to_lvar l, \') |l \'. (l, \') \ sem C S}" then obtain l where "\ = recover_p_state to_pval to_lvar l" "(l, \') \ sem C S" by blast then obtain x where "(l, x) \ S" "\C, x\ \ \'" by (metis fst_conv in_sem snd_conv) then have "l = copy_p_state to_pvar to_lval x" using assms(3) by blast moreover have "\ = x" proof (rule ext) fix y show "\ y = x y" by (simp add: \\ = recover_p_state to_pval to_lvar l\ assms(1) assms(2) calculation copy_p_state_def recover_p_state_def) qed ultimately show "(\, \') \ set_of_traces C" by (simp add: \\C, x\ \ \'\ set_of_traces_def) qed show "?B \ ?A" proof (rule subsetPairI) fix \ \' assume asm0: "(\, \') \ set_of_traces C" let ?l = "copy_p_state to_pvar to_lval \" have "(?l, \) \ S" using assms(3) by blast then have "(?l, \') \ sem C S" using asm0 in_set_of_traces_then_in_sem by blast moreover have "recover_p_state to_pval to_lvar ?l = \" proof (rule ext) fix x show "recover_p_state to_pval to_lvar (copy_p_state to_pvar to_lval \) x = \ x" by (simp add: assms(1) assms(2) copy_p_state_def recover_p_state_def) qed ultimately show "(\, \') \ {(recover_p_state to_pval to_lvar l, \') |l \'. (l, \') \ sem C S}" by force qed qed text \Theorem 3\ theorem proving_hyperproperties: fixes to_lvar :: "'pvar \ 'lvar" fixes to_lval :: "'pval \ 'lval" assumes "injective to_lvar" and "injective to_lval" shows "\P Q::('lvar, 'lval, 'pvar, 'pval) state hyperassertion. (\C. hypersat C H \ \ {P} C {Q})" proof - obtain to_pval :: "'lval \ 'pval" where r1: "\x. to_pval (to_lval x) = x" using assms(2) injective_then_exists_inverse by blast obtain to_pvar :: "'lvar \ 'pvar" where r2: "\x. to_pvar (to_lvar x) = x" using assms(1) injective_then_exists_inverse by blast let ?P = "\S. S = {(copy_p_state to_pvar to_lval \, \) |\. True }" let ?Q = "\S. H { (recover_p_state to_pval to_lvar l, \') |l \'. (l, \') \ S }" have "\C. hypersat C H \ \ {?P} C {?Q}" proof fix C assume "hypersat C H" show "\ {?P} C {?Q}" proof (rule hyper_hoare_tripleI) fix S assume "S = {(copy_p_state to_pvar to_lval \, \) |\. True}" have "{(recover_p_state to_pval to_lvar l, \') |l \'. (l, \') \ sem C S} = set_of_traces C" using \S = {(copy_p_state to_pvar to_lval \, \) |\. True}\ set_of_traces_same[of to_pvar to_lvar to_pval to_lval] r1 r2 by presburger then show "H {(recover_p_state to_pval to_lvar l, \') |l \'. (l, \') \ sem C S}" using \hypersat C H\ hypersat_def by metis qed next fix C let ?S = "{(copy_p_state to_pvar to_lval \, \) |\. True }" assume "\ {?P} C {?Q}" then have "?Q (sem C ?S)" by (simp add: hyper_hoare_triple_def) moreover have "{(recover_p_state to_pval to_lvar l, \') |l \'. (l, \') \ sem C ?S} = set_of_traces C" using r1 r2 set_of_traces_same[of to_pvar to_lvar to_pval to_lval] by presburger ultimately show "hypersat C H" by (simp add: hypersat_def) qed then show ?thesis by auto qed text \Hypersafety, hyperliveness\ definition max_k where "max_k k S \ finite S \ card S \ k" definition hypersafety where "hypersafety P \ (\S. \ P S \ (\S'. S \ S' \ \ P S'))" definition k_hypersafety where "k_hypersafety k P \ (\S. \ P S \ (\S'. S' \ S \ max_k k S' \ (\S''. S' \ S'' \ \ P S'')))" definition hyperliveness where "hyperliveness P \ (\S. \S'. S \ S' \ P S')" lemma k_hypersafetyI: assumes "\S. \ P S \ \S'. S' \ S \ max_k k S' \ (\S''. S' \ S'' \ \ P S'')" shows "k_hypersafety k P" by (simp add: assms k_hypersafety_def) lemma hypersafetyI: assumes "\S S'. \ P S \ S \ S' \ \ P S'" shows "hypersafety P" by (metis assms hypersafety_def) lemma hyperlivenessI: assumes "\S. \S'. S \ S' \ P S'" shows "hyperliveness P" using assms hyperliveness_def by blast lemma k_hypersafe_is_hypersafe: assumes "k_hypersafety k P" shows "hypersafety P" by (metis (full_types) assms dual_order.trans hypersafety_def k_hypersafety_def) lemma one_safety_equiv: assumes "sat H" shows "k_hypersafety 1 H \ (\P. \S. H S \ (\\ \ S. P \))" (is "?A \ ?B") proof assume ?B then obtain P where asm0: "\S. H S \ (\\ \ S. P \)" by auto show ?A proof (rule k_hypersafetyI) fix S assume asm1: "\ H S" then obtain \ where "\ \ S" "\ P \" using asm0 by blast let ?S = "{\}" have "?S \ S \ max_k 1 ?S \ (\S''. ?S \ S'' \ \ H S'')" using \\ P \\ \\ \ S\ asm0 max_k_def by fastforce then show "\S'\S. max_k 1 S' \ (\S''. S' \ S'' \ \ H S'')" by blast qed next assume ?A let ?P = "\\. H {\}" have "\S. H S \ (\\ \ S. ?P \)" proof fix S assume "H S" then show "\\ \ S. ?P \" using \k_hypersafety 1 H\ hypersafety_def k_hypersafe_is_hypersafe by auto next fix S assume asm0: "\\ \ S. ?P \" show "H S" proof (rule ccontr) assume "\ H S" then obtain S' where "S' \ S \ max_k 1 S' \ (\S''. S' \ S'' \ \ H S'')" by (metis \k_hypersafety 1 H\ k_hypersafety_def) then show False proof (cases "S' = {}") case True then show ?thesis by (metis \S' \ S \ max_k 1 S' \ (\S''. S' \ S'' \ \ H S'')\ assms empty_subsetI sat_def) next case False then obtain \ where "\ \ S'" by blast then have "card S' = 1" by (metis False One_nat_def Suc_leI \S' \ S \ max_k 1 S' \ (\S''. S' \ S'' \ \ H S'')\ card_gt_0_iff le_antisym max_k_def) then have "S' = {\}" using \\ \ S'\ card_1_singletonE by auto then show ?thesis using \S' \ S \ max_k 1 S' \ (\S''. S' \ S'' \ \ H S'')\ asm0 by fastforce qed qed qed then show ?B by blast qed definition hoarify where "hoarify P Q S \ (\p \ S. fst p \ P \ snd p \ Q)" lemma hoarify_hypersafety: "hypersafety (hoarify P Q)" by (metis (no_types, opaque_lifting) hoarify_def hypersafetyI subsetD) theorem hypersafety_1_hoare_logic: "k_hypersafety 1 (hoarify P Q)" proof (rule k_hypersafetyI) fix S assume "\ hoarify P Q S" then obtain \ where "\ \ S" "fst \ \ P" "snd \ \ Q" using hoarify_def by blast let ?S = "{\}" have "?S \ S \ max_k 1 ?S \ (\S''. ?S \ S'' \ \ hoarify P Q S'')" by (metis Compl_iff One_nat_def \\ \ S\ \fst \ \ P\ \snd \ \ Q\ card.empty card.insert compl_le_compl_iff empty_not_insert finite.intros(1) finite.intros(2) hoarify_def insert_absorb le_numeral_extra(4) max_k_def subset_Compl_singleton) then show "\S'\S. max_k 1 S' \ (\S''. S' \ S'' \ \ hoarify P Q S'')" by meson qed definition incorrectnessify where "incorrectnessify P Q S \ (\\' \ Q. \\ \ P. (\, \') \ S)" lemma incorrectnessify_liveness: assumes "P \ {}" shows "hyperliveness (incorrectnessify P Q)" proof (rule hyperlivenessI) fix S obtain \ where asm0: "\ \ P" using assms by blast let ?S = "S \ {(\, \') |\'. \' \ Q}" have "incorrectnessify P Q ?S" using asm0 incorrectnessify_def by force then show "\S'. S \ S' \ incorrectnessify P Q S'" using sup.cobounded1 by blast qed definition real_incorrectnessify where "real_incorrectnessify P Q S \ (\\ \ P. \\' \ Q. (\, \') \ S)" lemma real_incorrectnessify_liveness: assumes "Q \ {}" shows "hyperliveness (real_incorrectnessify P Q)" by (metis UNIV_I assms equals0I hyperliveness_def real_incorrectnessify_def subsetI) text \Verifying GNI\ definition gni_hyperassertion :: "'n \ 'n \ ('n \ 'v) hyperassertion" where "gni_hyperassertion h l S \ (\\ \ S. \v. \\' \ S. \' h = v \ \ l = \' l)" definition semify where "semify \ S = { (l, \') |\' \ l. (l, \) \ S \ (\, \') \ \ }" definition hyperprop_hht where "hyperprop_hht P Q \ \ (\S. P S \ Q (semify \ S))" -text \Footnote 4\ +text \Theorem 4\ theorem any_hht_hyperprop: "\ {P} C {Q} \ hypersat C (hyperprop_hht P Q)" (is "?A \ ?B") proof have "\S. semify (set_of_traces C) S = sem C S" proof - fix S have "\l \'. (l, \') \ sem C S \ (l, \') \ semify (set_of_traces C) S" proof - fix l \' have "(l, \') \ sem C S \ (\\. (l, \) \ S \ \C, \\ \ \')" by (simp add: in_sem) also have "... \ (\\. (l, \) \ S \ (\, \') \ set_of_traces C)" using set_of_traces_def by fastforce then show "(l, \') \ sem C S \ (l, \') \ semify (set_of_traces C) S" by (simp add: calculation semify_def) qed then show "semify (set_of_traces C) S = sem C S" by auto qed show "?A \ ?B" by (simp add: \\S. semify (set_of_traces C) S = sem C S\ hyper_hoare_tripleE hyperprop_hht_def hypersat_def) show "?B \ ?A" by (simp add: \\S. semify (set_of_traces C) S = sem C S\ hyper_hoare_triple_def hyperprop_hht_def hypersat_def) qed end \ No newline at end of file diff --git a/thys/HyperHoareLogic/ROOT b/thys/HyperHoareLogic/ROOT --- a/thys/HyperHoareLogic/ROOT +++ b/thys/HyperHoareLogic/ROOT @@ -1,13 +1,18 @@ chapter AFP session HyperHoareLogic = HOL + options [timeout = 600] theories - Language - Logic - Examples - ProgramHyperproperties - Expressivity + PaperResults + Language + Logic + ProgramHyperproperties + SyntacticAssertions + Loops + Expressivity + Compositionality + ExamplesCompositionality + TotalLogic document_files "root.bib" "root.tex" diff --git a/thys/HyperHoareLogic/SyntacticAssertions.thy b/thys/HyperHoareLogic/SyntacticAssertions.thy new file mode 100644 --- /dev/null +++ b/thys/HyperHoareLogic/SyntacticAssertions.thy @@ -0,0 +1,1656 @@ +section \Syntactic Assertions\ + +theory SyntacticAssertions + imports Logic Loops ProgramHyperproperties Compositionality +begin + +subsection \Preliminaries: Types, expressions, 'a assertions\ + +type_synonym var = nat +type_synonym qstate = nat +type_synonym qvar = nat + +type_synonym 'a nstate = "(var, 'a, var, 'a) state" +type_synonym 'a npstate = "(var, 'a) pstate" + +type_synonym 'a binop = "'a \ 'a \ 'a" +type_synonym 'a comp = "'a \ 'a \ bool" + +text \Quantified variables and quantified states are represented as de Bruijn indices (natural numbers).\ + +datatype 'a exp = + EPVar qstate var \\\\\<^sup>P(x)\: Program variable\ + | ELVar qstate var \\\\\<^sup>L(x)\: Logical variable\ + | EQVar qvar \\\y\: Quantified variable\ + | EConst 'a + | EBinop "'a exp" "'a binop" "'a exp" \\\e \ e\\ + | EFun "'a \ 'a" "'a exp" \\\f(e)\\ + +text \Quantified variables and quantified states are represented as de Bruijn indices (natural numbers). +Thus, quantifiers do not have a name for the variable or state they quantify over.\ + +datatype 'a assertion = + AConst bool + | AComp "'a exp" "'a comp" "'a exp" \\\e \ e\\ + | AForallState "'a assertion" \\\\<\>. A\\ + | AExistsState "'a assertion" \\\\<\>. A\\ + | AForall "'a assertion" \\\\y. A\\ + | AExists "'a assertion" \\\\y. A\\ + | AOr "'a assertion" "'a assertion" \\\A \ A\\ + | AAnd "'a assertion" "'a assertion" \\\A \ A\\ + +text \We use a list of values and a list of states to track quantified values and states, respectively.\ + +fun interp_exp :: "'a list \ 'a nstate list \ 'a exp \ 'a" where + "interp_exp vals states (EPVar st x) = snd (states ! st) x" +| "interp_exp vals states (ELVar st x) = fst (states ! st) x" +| "interp_exp vals states (EQVar x) = vals ! x" +| "interp_exp vals states (EConst v) = v" +| "interp_exp vals states (EBinop e1 op e2) = op (interp_exp vals states e1) (interp_exp vals states e2)" +| "interp_exp vals states (EFun f e) = f (interp_exp vals states e)" + +fun sat_assertion :: "'a list \ 'a nstate list \ 'a assertion \ 'a nstate set \ bool" where + "sat_assertion vals states (AConst b) _ \ b" +| "sat_assertion vals states (AComp e1 cmp e2) _ \ cmp (interp_exp vals states e1) (interp_exp vals states e2)" +| "sat_assertion vals states (AForallState A) S \ (\\\S. sat_assertion vals (\ # states) A S)" +| "sat_assertion vals states (AExistsState A) S \ (\\\S. sat_assertion vals (\ # states) A S)" +| "sat_assertion vals states (AForall A) S \ (\v. sat_assertion (v # vals) states A S)" +| "sat_assertion vals states (AExists A) S \ (\v. sat_assertion (v # vals) states A S)" +| "sat_assertion vals states (AAnd A B) S \ (sat_assertion vals states A S \ sat_assertion vals states B S)" +| "sat_assertion vals states (AOr A B) S \ (sat_assertion vals states A S \ sat_assertion vals states B S)" + +text \Negation and implication are defined on top of this base language.\ + +definition neg_cmp :: "'a comp \ 'a comp" where + "neg_cmp cmp v1 v2 \ \ (cmp v1 v2)" + +fun ANot where + "ANot (AConst b) = AConst (\ b)" +| "ANot (AComp e1 cmp e2) = AComp e1 (neg_cmp cmp) e2" +| "ANot (AForallState A) = AExistsState (ANot A)" +| "ANot (AExistsState A) = AForallState (ANot A)" +| "ANot (AOr A B) = AAnd (ANot A) (ANot B)" +| "ANot (AAnd A B) = AOr (ANot A) (ANot B)" +| "ANot (AForall A) = AExists (ANot A)" +| "ANot (AExists A) = AForall (ANot A)" + +definition AImp where + "AImp A B = AOr (ANot A) B" + +lemma sat_assertion_Not: + "sat_assertion vals states (ANot A) S \ \(sat_assertion vals states A S)" + by (induct A arbitrary: vals states) (simp_all add: neg_cmp_def) + +lemma sat_assertion_Imp: + "sat_assertion vals states (AImp A B) S \ (sat_assertion vals states A S \ sat_assertion vals states B S)" + by (simp add: AImp_def sat_assertion_Not) + +abbreviation interp_assert where "interp_assert \ sat_assertion [] []" + + +subsection \Assume rule\ + +fun transform_assume :: "'a assertion \ 'a assertion \ 'a assertion" where + "transform_assume _ (AConst b) = AConst b" +| "transform_assume _ (AComp e1 cmp e2) = AComp e1 cmp e2" +| "transform_assume b (AForallState A) = AForallState (AImp b (transform_assume b A))" +| "transform_assume b (AExistsState A) = AExistsState (AAnd b (transform_assume b A))" +| "transform_assume b (AForall A) = AForall (transform_assume b A)" +| "transform_assume b (AExists A) = AExists (transform_assume b A)" +| "transform_assume b (AAnd A B) = AAnd (transform_assume b A) (transform_assume b B)" +| "transform_assume b (AOr A B) = AOr (transform_assume b A) (transform_assume b B)" + + +definition same_syn_sem :: "'a assertion \ ('a npstate \ bool) \ bool" + where + "same_syn_sem bsyn bsem \ + (\states vals S. length states > 0 \ bsem (snd (hd states)) = sat_assertion vals states bsyn S)" + +lemma same_syn_semI: + assumes "\states vals S. length states > 0 \ bsem (snd (hd states)) \ sat_assertion vals states bsyn S" + shows "same_syn_sem bsyn bsem" + by (simp add: assms same_syn_sem_def) + +lemma transform_assume_valid: + assumes "same_syn_sem bsyn bsem" + shows "sat_assertion vals states A (Set.filter (bsem \ snd) S) + \ sat_assertion vals states (transform_assume bsyn A) S" +proof (induct A arbitrary: vals states) + case (AForallState A) + let ?S = "Set.filter (bsem \ snd) S" + let ?A = "transform_assume bsyn A" + have "sat_assertion vals states (AForallState A) ?S \ (\\\?S. sat_assertion vals (\ # states) A ?S)" + by force + also have "... \ (\\\?S. sat_assertion vals (\ # states) ?A S)" + using AForallState by presburger + also have "... \ (\\\S. bsem (snd \) \ sat_assertion vals (\ # states) ?A S)" + by fastforce + also have "... \ (\\\S. sat_assertion vals (\ # states) bsyn S \ sat_assertion vals (\ # states) ?A S)" + using assms same_syn_sem_def[of bsyn bsem] by auto + also have "... \ (\\\S. sat_assertion vals (\ # states) (AImp bsyn ?A) S)" + using sat_assertion_Imp by fast + also have "... \ sat_assertion vals states (AForallState (AImp bsyn ?A)) S" + using sat_assertion.simps(2) by force + finally show ?case + using transform_assume.simps(1) by fastforce +next + case (AExistsState A) + let ?S = "Set.filter (bsem \ snd) S" + let ?A = "transform_assume bsyn A" + have "sat_assertion vals states (AExistsState A) ?S \ (\\\?S. sat_assertion vals (\ # states) A ?S)" + by force + also have "... \ (\\\?S. sat_assertion vals (\ # states) ?A S)" + using AExistsState by presburger + also have "... \ (\\\S. bsem (snd \) \ sat_assertion vals (\ # states) ?A S)" + by force + also have "... \ (\\\S. sat_assertion vals (\ # states) bsyn S \ sat_assertion vals (\ # states) ?A S)" + using assms same_syn_sem_def[of bsyn bsem] by auto + also have "... \ (\\\S. sat_assertion vals (\ # states) (AAnd bsyn ?A) S)" + by simp + also have "... \ sat_assertion vals states (AExistsState (AAnd bsyn ?A)) S" + using sat_assertion.simps(3) by force + then show ?case + using calculation transform_assume.simps(2) by fastforce +next + case (AForall A) + let ?S = "Set.filter (bsem \ snd) S" + let ?A = "transform_assume bsyn A" + have "sat_assertion vals states (AForall A) ?S \ (\v. sat_assertion (v # vals) states A ?S)" + by force + also have "... \ (\v. sat_assertion (v # vals) states ?A S)" + using AForall by presburger + also have "... \ sat_assertion vals states (AForall ?A) S" + by simp + then show ?case + using calculation transform_assume.simps(3) by fastforce +next + case (AExists A) + let ?S = "Set.filter (bsem \ snd) S" + let ?A = "transform_assume bsyn A" + have "sat_assertion vals states (AExists A) ?S \ (\v. sat_assertion (v # vals) states A ?S)" + by force + also have "... \ (\v. sat_assertion (v # vals) states ?A S)" + using AExists by presburger + also have "... \ sat_assertion vals states (AExists ?A) S" + by simp + then show ?case + using calculation transform_assume.simps(4) by fastforce +qed (simp_all) + + + +subsubsection \Program expressions (values)\ + +datatype 'a pexp = + PVar var \\Normal variable, like x\ + | PConst 'a + | PBinop "'a pexp" "'a binop" "'a pexp" + | PFun "'a \ 'a" "'a pexp" + + +fun interp_pexp :: "'a pexp \ 'a npstate \ 'a" + where + "interp_pexp (PVar x) \ = \ x" +| "interp_pexp (PConst n) _ = n" +| "interp_pexp (PBinop p1 op p2) \ = op (interp_pexp p1 \) (interp_pexp p2 \)" +| "interp_pexp (PFun f p) \ = f (interp_pexp p \)" + +fun pexp_to_exp where + "pexp_to_exp st (PVar x) = EPVar st x" +| "pexp_to_exp _ (PConst n) = EConst n" +| "pexp_to_exp st (PBinop p1 op p2) = EBinop (pexp_to_exp st p1) op (pexp_to_exp st p2)" +| "pexp_to_exp st (PFun f p) = EFun f (pexp_to_exp st p)" + +lemma same_syn_sem_exp: + "interp_pexp p (snd (states ! st)) = interp_exp vals states (pexp_to_exp st p)" +proof (induct p) + case (PVar x) + then show ?case + using hd_conv_nth by force +qed (simp_all) + + +subsubsection \Program expressions (booleans)\ + +datatype 'a pbexp = + PBConst bool + | PBAnd "'a pbexp" "'a pbexp" + | PBOr "'a pbexp" "'a pbexp" + | PBComp "'a pexp" "'a comp" "'a pexp" + +fun interp_pbexp :: "'a pbexp \ 'a npstate \ bool" + where + "interp_pbexp (PBConst b) _ \ b" +| "interp_pbexp (PBAnd pb1 pb2) \ \ interp_pbexp pb1 \ \ interp_pbexp pb2 \" +| "interp_pbexp (PBOr pb1 pb2) \ \ interp_pbexp pb1 \ \ interp_pbexp pb2 \" +| "interp_pbexp (PBComp p1 cmp p2) \ \ cmp (interp_pexp p1 \) (interp_pexp p2 \)" + +fun pbexp_to_assertion where + "pbexp_to_assertion _ (PBConst b) = AConst b" +| "pbexp_to_assertion st (PBAnd pb1 pb2) = AAnd (pbexp_to_assertion st pb1) (pbexp_to_assertion st pb2)" +| "pbexp_to_assertion st (PBOr pb1 pb2) = AOr (pbexp_to_assertion st pb1) (pbexp_to_assertion st pb2)" +| "pbexp_to_assertion st (PBComp p1 cmp p2) = AComp (pexp_to_exp st p1) cmp (pexp_to_exp st p2)" + +lemma same_syn_sem_assertion: + "interp_pbexp pb (snd (states ! st)) = sat_assertion vals states (pbexp_to_assertion st pb) S" +proof (induct pb) + case (PBComp x1 x2 x3) + then show ?case + by (metis interp_pbexp.simps(4) pbexp_to_assertion.simps(4) same_syn_sem_exp sat_assertion.simps(2)) +qed (simp_all) + +lemma pexp_to_exp_same: + shows "same_syn_sem (pbexp_to_assertion 0 pb) (interp_pbexp pb)" +proof (rule same_syn_semI) + fix states :: "'a nstate list" + fix vals S + assume "0 < length states" + then have "sat_assertion vals states (pbexp_to_assertion 0 pb) S = sat_assertion [] states (pbexp_to_assertion 0 pb) S" + using same_syn_sem_assertion by blast + then show "interp_pbexp pb (snd (hd states)) = sat_assertion vals states (pbexp_to_assertion 0 pb) S" + by (metis \0 < length states\ hd_conv_nth length_greater_0_conv same_syn_sem_assertion) +qed + + + +subsubsection \Syntactic rule for assume\ + +theorem rule_assume_syntactic_general: + "\ { sat_assertion states vals (transform_assume (pbexp_to_assertion 0 pb) P) } Assume (interp_pbexp pb) {sat_assertion states vals P}" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "sat_assertion states vals (transform_assume (pbexp_to_assertion 0 pb) P) S" + then have "sat_assertion states vals P (Set.filter (interp_pbexp pb \ snd) S)" + using pexp_to_exp_same transform_assume_valid by blast + then show "sat_assertion states vals P (sem (Assume (interp_pbexp pb)) S)" + by (simp add: assume_sem) +qed + +theorem rule_assume_syntactic: + "\ { interp_assert (transform_assume (pbexp_to_assertion 0 pb) P) } Assume (interp_pbexp pb) {interp_assert P}" + by (simp add: rule_assume_syntactic_general) + + + +subsection \Havoc rule\ + +subsubsection \Shifting variables\ + +fun insert_at where + "insert_at 0 x l = x # l" +| "insert_at (Suc n) x (t # q) = t # (insert_at n x q)" +| "insert_at (Suc n) x [] = [x]" + +lemma length_insert_at: + "length (insert_at n x l) = length l + 1" +proof (induct n arbitrary: l) + case (Suc n) + then show ?case + proof (cases l) + case (Cons t q) + then show ?thesis + by (simp add: Suc) + qed (simp) +qed (simp_all) + + +lemma insert_at_charact_1: + "n \ length l \ k < n \ (insert_at n x l) ! k = l ! k" +proof (induct k arbitrary: l n) + case 0 + then show ?case + by (metis bot_nat_0.not_eq_extremum insert_at.elims le_zero_eq list.size(3) nth_Cons_0) +next + case (Suc k) + then obtain nn t q where "n = Suc nn" "l = t # q" + by (metis Suc_less_eq2 le_antisym list.exhaust list.size(3) not_less_zero zero_le) + then show ?case + using Suc.hyps Suc.prems(1) Suc.prems(2) by force +qed + +lemma insert_at_charact_2: + "n \ length l \ (insert_at n x l) ! n = x" +proof (induct n arbitrary: l) + case (Suc n) + then show ?case + by (metis Suc_le_length_iff insert_at.simps(2) nth_Cons_Suc) +qed (simp) + +lemma insert_at_charact_3: + "n \ length l \ k \ n \ (insert_at n x l) ! (Suc k) = l ! k" +proof (induct n arbitrary: l k) + case (Suc xa) + then obtain t q kk where "k = Suc kk" "l = t # q" + by (meson Suc_le_D Suc_le_length_iff) + then show ?case + using Suc.hyps Suc.prems(1) Suc.prems(2) by auto +qed (simp) + + +(* Shift only stuff above *) +fun shift_vars_exp where + "shift_vars_exp n (EQVar x) = (if x \ n then EQVar (Suc x) else EQVar x)" +| "shift_vars_exp n (EBinop e1 op e2) = EBinop (shift_vars_exp n e1) op (shift_vars_exp n e2)" +| "shift_vars_exp n (EFun p e) = EFun p (shift_vars_exp n e)" +| "shift_vars_exp _ e = e" + +fun shift_states_exp where + "shift_states_exp n (EPVar \ x) = (if \ \ n then EPVar (Suc \) x else EPVar \ x)" +| "shift_states_exp n (ELVar \ x) = (if \ \ n then ELVar (Suc \) x else ELVar \ x)" +| "shift_states_exp n (EBinop e1 op e2) = EBinop (shift_states_exp n e1) op (shift_states_exp n e2)" +| "shift_states_exp n (EFun p e) = EFun p (shift_states_exp n e)" +| "shift_states_exp _ e = e" + +fun wf_exp :: "nat \ nat \ 'a exp \ bool" where + "wf_exp nv ns (EPVar st _) \ st < ns" +| "wf_exp nv ns (ELVar st _) \ st < ns" +| "wf_exp nv ns (EQVar x) \ x < nv" +| "wf_exp nv ns (EBinop e1 _ e2) \ wf_exp nv ns e1 \ wf_exp nv ns e2" +| "wf_exp nv ns (EFun f e) \ wf_exp nv ns e" +| "wf_exp nv ns (EConst _) \ True" + +lemma wf_shift_vars_exp: + assumes "wf_exp nv ns e" + shows "wf_exp (Suc nv) ns (shift_vars_exp n e)" + using assms + by (induct e) simp_all + +lemma wf_shift_states_exp: + assumes "wf_exp nv ns e" + shows "wf_exp nv (Suc ns) (shift_states_exp n e)" + using assms + by (induct e) simp_all + +lemma shift_vars_exp_charact: + assumes "n \ length vals" + shows "interp_exp vals states e = interp_exp (insert_at n v vals) states (shift_vars_exp n e)" + using assms +proof (induct e) + case (EQVar x) + then show ?case + by (simp add: insert_at_charact_1 insert_at_charact_3) +qed (simp_all) + +lemma shift_states_exp_charact: + assumes "n \ length states" + shows "interp_exp vals states e = interp_exp vals (insert_at n \ states) (shift_states_exp n e)" + using assms +proof (induct e) + case (EPVar x1 x2) + then show ?case + by (simp add: insert_at_charact_1 insert_at_charact_3) +next + case (ELVar x1 x2) + then show ?case + by (simp add: insert_at_charact_1 insert_at_charact_3) +qed (simp_all) + + +fun shift_vars where + "shift_vars n (AConst b) = AConst b" +| "shift_vars n (AComp e1 cmp e2) = AComp (shift_vars_exp n e1) cmp (shift_vars_exp n e2)" +| "shift_vars n (AForall A) = AForall (shift_vars (Suc n) A)" +| "shift_vars n (AExists A) = AExists (shift_vars (Suc n) A)" +| "shift_vars n (AForallState A) = AForallState (shift_vars n A)" +| "shift_vars n (AExistsState A) = AExistsState (shift_vars n A)" +| "shift_vars n (AOr A B) = AOr (shift_vars n A) (shift_vars n B)" +| "shift_vars n (AAnd A B) = AAnd (shift_vars n A) (shift_vars n B)" + +lemma shift_vars_charact: + assumes "n \ length vals" + shows "sat_assertion vals states A S \ sat_assertion (insert_at n x vals) states (shift_vars n A) S" + using assms +proof (induct A arbitrary: vals states n) + case (AComp x1a x2 x3a) + then show ?case + using shift_vars_exp_charact by fastforce +next + case (AForall A) + have "sat_assertion vals states (AForall A) S \ (\v. sat_assertion (v # vals) states A S)" + by simp + also have "... \ (\v. sat_assertion (insert_at (Suc n) x (v # vals)) states (shift_vars (Suc n) A) S)" + using AForall(2) AForall(1)[of "Suc n" _ states] by simp + also have "... \ (\v. sat_assertion (v # insert_at n x vals) states (shift_vars (Suc n) A) S)" + by simp + also have "... \ sat_assertion (insert_at n x vals) states (AForall (shift_vars (Suc n) A)) S" + by simp + then show "sat_assertion vals states (AForall A) S = sat_assertion (insert_at n x vals) states (shift_vars n (AForall A)) S" + using calculation by simp +next + case (AExists A) + have "sat_assertion vals states (AExists A) S \ (\v. sat_assertion (v # vals) states A S)" + by simp + also have "... \ (\v. sat_assertion (insert_at (Suc n) x (v # vals)) states (shift_vars (Suc n) A) S)" + using AExists(2) AExists(1)[of "Suc n" _ states] by simp + also have "... \ (\v. sat_assertion (v # insert_at n x vals) states (shift_vars (Suc n) A) S)" + by simp + also have "... \ sat_assertion (insert_at n x vals) states (AExists (shift_vars (Suc n) A)) S" + by simp + then show "sat_assertion vals states (AExists A) S = sat_assertion (insert_at n x vals) states (shift_vars n (AExists A)) S" + using calculation by simp +qed (simp_all) + + + +subsubsection \Expressions (Boolean and values)\ + +definition update_state where + "update_state \ x v = (fst \, (snd \)(x := v))" + + +(* Replacing \(x) by v *) +(* for havoc, should be used with (EQVar v) *) +fun subst_exp_single :: "qstate \ var \ 'a exp \ 'a exp \ 'a exp" where + "subst_exp_single \ x e' (EPVar st y) = (if \ = st \ x = y then e' else EPVar st y)" +| "subst_exp_single \ x e' (EBinop e1 bop e2) = EBinop (subst_exp_single \ x e' e1) bop (subst_exp_single \ x e' e2)" +| "subst_exp_single \ x e' (EFun f e) = EFun f (subst_exp_single \ x e' e)" +| "subst_exp_single _ _ _ e = e" (* Logical variables, quantified variables, constants *) + +lemma wf_subst_exp: + assumes "wf_exp nv ns e" + and "wf_exp nv ns e'" + shows "wf_exp nv ns (subst_exp_single \ x e' e)" + using assms + by (induct e) simp_all + + + +lemma subst_exp_single_charact: + assumes "interp_exp vals states e' = snd (states ! st) x" + shows "interp_exp vals states (subst_exp_single st x e' e) = interp_exp vals states e" + using assms + by (induct e) simp_all + + +definition subst_state where + "subst_state x pe \ = (fst \, (snd \)(x := interp_pexp pe (snd \)))" + +definition update_state_at where + "update_state_at states n x v = list_update states n (update_state (states ! n) x v)" + +lemma update_state_at_fst: + "fst (update_state_at states n x v ! st) = fst (states ! st)" +proof (cases "n = st") + case True + then show ?thesis + by (metis fst_conv linorder_not_less list_update_beyond nth_list_update_eq update_state_at_def update_state_def) +next + case False + then show ?thesis + by (simp add: update_state_at_def) +qed + +lemma update_state_at_snd_1: + "x \ y \ snd (update_state_at states n x v ! st) y = snd (states ! st) y" + apply (cases "n = st") + apply (metis fun_upd_other linorder_not_less list_update_beyond nth_list_update_eq snd_conv update_state_at_def update_state_def) + by (simp add: update_state_at_def) + +lemma update_state_at_snd_2: + "st \ n \ snd (update_state_at states n x v ! st) y = snd (states ! st) y" + by (simp add: update_state_at_def) + +lemma update_state_at_snd_3: + assumes "n < length states" + shows "snd (update_state_at states n x v ! n) x = v" + by (simp add: assms update_state_at_def update_state_def) + +lemma subst_exp_more_complex_charact: + assumes "states' = update_state_at states st x (interp_exp vals states e')" + and "st < length states" + shows "interp_exp vals states (subst_exp_single st x e' e) = interp_exp vals states' e" + using assms +proof (induct e) + case (EPVar \ y) + then show ?case + by (metis interp_exp.simps(1) subst_exp_single.simps(1) update_state_at_snd_1 update_state_at_snd_2 update_state_at_snd_3) +next + case (ELVar x1 x2) + then show ?case + by (simp add: update_state_at_fst) +qed (simp_all) + + +subsubsection \Assertions\ + +fun subst_assertion_single :: "qstate \ var \ 'a exp \ 'a assertion \ 'a assertion" where + "subst_assertion_single st x e (AConst b) = AConst b" +| "subst_assertion_single st x e (AComp e1 cmp e2) = AComp (subst_exp_single st x e e1) cmp (subst_exp_single st x e e2)" +| "subst_assertion_single st x e (AForall A) = AForall (subst_assertion_single st x (shift_vars_exp 0 e) A)" +| "subst_assertion_single st x e (AExists A) = AExists (subst_assertion_single st x (shift_vars_exp 0 e) A)" +| "subst_assertion_single st x e (AOr A B) = AOr (subst_assertion_single st x e A) (subst_assertion_single st x e B)" +| "subst_assertion_single st x e (AAnd A B) = AAnd (subst_assertion_single st x e A) (subst_assertion_single st x e B)" +| "subst_assertion_single st x e (AForallState A) = AForallState (subst_assertion_single (Suc st) x (shift_states_exp 0 e) A)" +| "subst_assertion_single st x e (AExistsState A) = AExistsState (subst_assertion_single (Suc st) x (shift_states_exp 0 e) A)" + +fun wf_assertion_aux :: "nat \ nat \ 'a assertion \ bool" where + "wf_assertion_aux nv ns (AConst b) \ True" +| "wf_assertion_aux nv ns (AComp e1 cmp e2) \ wf_exp nv ns e1 \ wf_exp nv ns e2" +| "wf_assertion_aux nv ns (AAnd A B) \ wf_assertion_aux nv ns A \ wf_assertion_aux nv ns B" +| "wf_assertion_aux nv ns (AOr A B) \ wf_assertion_aux nv ns A \ wf_assertion_aux nv ns B" + +| "wf_assertion_aux nv ns (AForall A) \ wf_assertion_aux (Suc nv) ns A" +| "wf_assertion_aux nv ns (AExists A) \ wf_assertion_aux (Suc nv) ns A" +| "wf_assertion_aux nv ns (AForallState A) \ wf_assertion_aux nv (Suc ns) A" +| "wf_assertion_aux nv ns (AExistsState A) \ wf_assertion_aux nv (Suc ns) A" + + +abbreviation wf_assertion where "wf_assertion \ wf_assertion_aux 0 0" + + +lemma wf_shift_vars: + assumes "wf_assertion_aux nv ns A" + shows "wf_assertion_aux (Suc nv) ns (shift_vars n A)" + using assms + by (induct A arbitrary: n nv ns) (simp_all add: wf_shift_vars_exp) + +lemma wf_subst_assertion: + assumes "wf_assertion_aux nv ns A" + and "wf_exp nv ns e" + shows "wf_assertion_aux nv ns (subst_assertion_single \ x e A)" + using assms +proof (induct A arbitrary: nv ns e \) + case (AComp x1a x2 x3a) + then show ?case + by (simp add: wf_subst_exp) +qed (simp_all add: wf_shift_vars_exp wf_shift_states_exp) + +lemma subst_assertion_single_charact: + assumes "interp_exp vals states e = snd (states ! st) x" + shows "sat_assertion vals states (subst_assertion_single st x e A) S \ sat_assertion vals states A S" + using assms +proof (induct A arbitrary: vals states st e) + case (AForallState A) + have "sat_assertion vals states (AForallState A) S \ (\\ \ S. sat_assertion vals (\ # states) A S)" + by simp + also have "... \ (\\ \ S. sat_assertion vals (\ # states) (subst_assertion_single (Suc st) x (shift_states_exp 0 e) A) S)" + using AForallState(1)[of vals _ "shift_states_exp 0 e" "Suc st"] AForallState(2) + by (metis insert_at.simps(1) nth_Cons_Suc shift_states_exp_charact zero_le) + finally show ?case by simp +next + case (AExistsState A) + have "sat_assertion vals states (AExistsState A) S \ (\\ \ S. sat_assertion vals (\ # states) A S)" + by simp + also have "... \ (\\ \ S. sat_assertion vals (\ # states) (subst_assertion_single (Suc st) x (shift_states_exp 0 e) A) S)" + by (metis AExistsState.hyps AExistsState.prems insert_at.simps(1) nth_Cons_Suc shift_states_exp_charact zero_le) + finally show ?case by simp +next + case (AForall A) + have "sat_assertion vals states (AForall A) S \ (\v. sat_assertion (v # vals) states A S)" + by simp + also have "... \ (\v. sat_assertion (v # vals) states (subst_assertion_single st x (shift_vars_exp 0 e) A) S)" + by (metis AForall.hyps AForall.prems insert_at.simps(1) shift_vars_exp_charact zero_le) + finally show ?case + by simp +next + case (AExists A) + have "sat_assertion vals states (AExists A) S \ (\v. sat_assertion (v # vals) states A S)" + by simp + also have "... \ (\v. sat_assertion (v # vals) states (subst_assertion_single st x (shift_vars_exp 0 e) A) S)" + by (metis AExists.hyps AExists.prems insert_at.simps(1) shift_vars_exp_charact zero_le) + finally show ?case + by simp +next + case (AComp e1 cmp e2) + then show ?case + using subst_exp_single_charact[of vals states e st x] by auto +qed (simp_all) + + + +lemma update_state_at_cons: + "update_state_at (\ # states) (Suc n) x v = \ # update_state_at states n x v" + by (simp add: update_state_at_def) + + +lemma subst_assertion_single_charact_better: + assumes "states' = update_state_at states st x (interp_exp vals states e)" + and "st < length states" + shows "sat_assertion vals states (subst_assertion_single st x e A) S \ sat_assertion vals states' A S" + using assms +proof (induct A arbitrary: vals states states' st e) + case (AComp x1a x2 x3a) + then show ?case + by (simp add: subst_exp_more_complex_charact) +next + case (AForallState A) + have "sat_assertion vals states (subst_assertion_single st x e (AForallState A)) S + \ (\\\S. sat_assertion vals (update_state_at (\ # states) (Suc st) x (interp_exp vals (\ # states) (shift_states_exp 0 e))) A S)" + by (simp add: AForallState.hyps AForallState.prems(2)) + also have "... \ (\\\S. sat_assertion vals (\ # update_state_at states st x (interp_exp vals states e)) A S)" + by (metis update_state_at_cons insert_at.simps(1) shift_states_exp_charact zero_le) + finally show "sat_assertion vals states (subst_assertion_single st x e (AForallState A)) S = sat_assertion vals states' (AForallState A) S" + by (simp add: AForallState.prems(1)) +next + case (AExistsState A) + have "sat_assertion vals states (subst_assertion_single st x e (AExistsState A)) S + \ (\\\S. sat_assertion vals (update_state_at (\ # states) (Suc st) x (interp_exp vals (\ # states) (shift_states_exp 0 e))) A S)" + by (simp add: AExistsState.hyps AExistsState.prems(2)) + also have "... \ (\\\S. sat_assertion vals (\ # update_state_at states st x (interp_exp vals states e)) A S)" + by (metis update_state_at_cons insert_at.simps(1) shift_states_exp_charact zero_le) + finally show "sat_assertion vals states (subst_assertion_single st x e (AExistsState A)) S = sat_assertion vals states' (AExistsState A) S" + by (simp add: AExistsState.prems(1)) +next + case (AForall A) + have "sat_assertion vals states (subst_assertion_single st x e (AForall A)) S + \ (\v. sat_assertion (v # vals) states (subst_assertion_single st x (shift_vars_exp 0 e) A) S)" + by (simp add: AForall.hyps AForall.prems(2)) + then show ?case + by (metis AForall.hyps AForall.prems(1) AForall.prems(2) insert_at.simps(1) sat_assertion.simps(5) shift_vars_exp_charact zero_le) +next + case (AExists A) + have "sat_assertion vals states (subst_assertion_single st x e (AExists A)) S + \ (\v. sat_assertion (v # vals) states (subst_assertion_single st x (shift_vars_exp 0 e) A) S)" + by (simp add: AExists.hyps AExists.prems(2)) + then show ?case + by (metis AExists.hyps AExists.prems(1) AExists.prems(2) insert_at.simps(1) sat_assertion.simps(6) shift_vars_exp_charact zero_le) +qed (simp_all) + + +subsubsection \Transformation for havoc\ + + +fun transform_havoc where + "transform_havoc x (AForallState A) = AForallState (AForall (subst_assertion_single 0 x (EQVar 0) (shift_vars 0 (transform_havoc x A))))" +| "transform_havoc x (AExistsState A) = AExistsState (AExists (subst_assertion_single 0 x (EQVar 0) (shift_vars 0 (transform_havoc x A))))" +| "transform_havoc x (AExists A) = AExists (transform_havoc x A)" +| "transform_havoc x (AForall A) = AForall (transform_havoc x A)" +| "transform_havoc x (AOr A B) = AOr (transform_havoc x A) (transform_havoc x B)" +| "transform_havoc x (AAnd A B) = AAnd (transform_havoc x A) (transform_havoc x B)" +| "transform_havoc x (AConst b) = AConst b" +| "transform_havoc x (AComp e1 cmp e2) = AComp e1 cmp e2" + + +lemma sem_havoc_bis: + "sem (Havoc x) S = {(fst \, (snd \)(x := v)) |\ v. \ \ S}" (is "?A = ?B") +proof + show "?B \ ?A" + using sem_havoc by fastforce + show "?A \ ?B" + proof + fix \ assume "\ \ ?A" + then obtain l \ v where "\ = (l, \(x := v))" "(l, \) \ S" + by (metis in_sem prod.collapse single_sem_Havoc_elim) + then show "\ \ ?B" + by auto + qed +qed + +lemma helper_update_state: + "(v # vals) ! 0 = snd ((update_state \ x v # states) ! 0) x" + by (simp add: update_state_def) + +lemma helper_S_update_states: + assumes "S' = { update_state \ x v |\ v. \ \ S}" + shows "(\\ \ S'. Q \) \ (\\ \ S. \v. Q (update_state \ x v))" +proof + show "\\\S'. Q \ \ \\\S. \v. Q (update_state \ x v)" + using assms by blast + show "\\\S. \v. Q (update_state \ x v) \ \\\S'. Q \" + using assms by force +qed + +lemma helper_S_update_states_exists: + assumes "S' = { update_state \ x v |\ v. \ \ S}" + shows "(\\ \ S'. Q \) \ (\\ \ S. \v. Q (update_state \ x v))" +proof + show "\\\S'. Q \ \ \\\S. \v. Q (update_state \ x v)" + using assms by force + show "\\\S. \v. Q (update_state \ x v) \ \\\S'. Q \" + using assms by blast +qed + + + +lemma equiv_havoc_transform: + assumes "S' = { update_state \ x v |\ v. \ \ S}" + shows "sat_assertion vals states P S' \ sat_assertion vals states (transform_havoc x P) S" +proof (induct P arbitrary: vals states) + case (AForallState P) + + let ?PP = "shift_vars 0 (transform_havoc x P)" + let ?P = "subst_assertion_single 0 x (EQVar 0) ?PP" + + have rr: "\\ v. sat_assertion (v # vals) (\ # states) ?P S + \ sat_assertion (v # vals) (update_state \ x v # states) ?P S" + proof - + fix \ v + have "sat_assertion (v # vals) (insert_at 0 \ states) (subst_assertion_single 0 x (EQVar 0) (shift_vars 0 (transform_havoc x P))) S = +sat_assertion (v # vals) (insert_at 0 (update_state \ x v) states) (subst_assertion_single 0 x (EQVar 0) (shift_vars 0 (transform_havoc x P))) S" + by (metis (no_types, lifting) One_nat_def helper_update_state insert_at.simps(1) interp_exp.simps(3) length_insert_at list_update_code(2) nth_Cons_0 subst_assertion_single_charact subst_assertion_single_charact_better trans_less_add2 update_state_at_def zero_less_Suc) + then show "sat_assertion (v # vals) (\ # states) ?P S \ sat_assertion (v # vals) (update_state \ x v # states) ?P S" + by simp + qed + have "sat_assertion vals states (transform_havoc x (AForallState P)) S \ sat_assertion vals states (AForallState (AForall ?P)) S" + by simp + also have "... \ (\\ \ S. \v. sat_assertion (v # vals) (update_state \ x v # states) ?P S)" + using rr by simp + also have "... \ (\\ \ S. \v. sat_assertion (v # vals) (update_state \ x v # states) ?PP S)" + using rr subst_assertion_single_charact[of _ _ _ _ x ?PP S] helper_update_state + by (metis interp_exp.simps(3)) + also have "... \ (\\ \ S. \v. sat_assertion vals (update_state \ x v # states) (transform_havoc x P) S)" + by (metis insert_at.simps(1) le0 shift_vars_charact) + also have "... \ (\\ \ S. \v. sat_assertion vals (update_state \ x v # states) P S')" + using AForallState.hyps AForallState.prems by force + also have "... \ sat_assertion vals states (AForallState P) S'" + using helper_S_update_states[of S' x S "\\. sat_assertion vals (\ # states) P S'"] assms by force + then show ?case + using calculation by blast +next + case (AExistsState P) + + let ?PP = "shift_vars 0 (transform_havoc x P)" + let ?P = "subst_assertion_single 0 x (EQVar 0) ?PP" + + have rr: "\\ v. sat_assertion (v # vals) (\ # states) ?P S + \ sat_assertion (v # vals) (update_state \ x v # states) ?P S" + proof - + fix \ v + have "sat_assertion (v # vals) (insert_at 0 \ states) (subst_assertion_single 0 x (EQVar 0) (shift_vars 0 (transform_havoc x P))) S = +sat_assertion (v # vals) (insert_at 0 (update_state \ x v) states) (subst_assertion_single 0 x (EQVar 0) (shift_vars 0 (transform_havoc x P))) S" + by (metis (no_types, lifting) One_nat_def helper_update_state insert_at.simps(1) insert_at_charact_2 interp_exp.simps(3) length_insert_at list_update_code(2) subst_assertion_single_charact subst_assertion_single_charact_better trans_less_add2 update_state_at_def zero_le zero_less_Suc) + then show "sat_assertion (v # vals) (\ # states) ?P S \ sat_assertion (v # vals) (update_state \ x v # states) ?P S" + by simp + qed + have "sat_assertion vals states (transform_havoc x (AExistsState P)) S \ sat_assertion vals states (AExistsState (AExists ?P)) S" + by simp + also have "... \ (\\ \ S. \v. sat_assertion (v # vals) (update_state \ x v # states) ?P S)" + using rr by simp + also have "... \ (\\ \ S. \v. sat_assertion (v # vals) (update_state \ x v # states) ?PP S)" + by (metis helper_update_state interp_exp.simps(3) subst_assertion_single_charact) + also have "... \ (\\ \ S. \v. sat_assertion vals (update_state \ x v # states) (transform_havoc x P) S)" + by (metis insert_at.simps(1) le0 shift_vars_charact) + also have "... \ (\\ \ S. \v. sat_assertion vals (update_state \ x v # states) P S')" + using AExistsState.hyps AExistsState.prems by force + also have "... \ sat_assertion vals states (AExistsState P) S'" + using helper_S_update_states_exists[of S' x S "\\. sat_assertion vals (\ # states) P S'"] assms + by simp + then show ?case + using calculation by blast +qed (simp_all) + + + + +subsubsection \Syntactic rule for havoc\ + +theorem rule_havoc_syntactic_general: + "\ { sat_assertion states vals (transform_havoc x P) } Havoc x {sat_assertion states vals P}" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "sat_assertion states vals (transform_havoc x P) S" + let ?S = "sem (Havoc x) S" + have "sat_assertion states vals P ?S \ sat_assertion states vals (transform_havoc x P) S" + proof (rule equiv_havoc_transform) + show "sem (Havoc x) S = {update_state \ x v |\ v. \ \ S}" + by (simp add: sem_havoc_bis update_state_def) + qed + then show "sat_assertion states vals P (sem (Havoc x) S)" + using asm0 by fastforce +qed + + +theorem rule_havoc_syntactic: + "\ { interp_assert (transform_havoc x P) } Havoc x {interp_assert P}" + by (simp add: rule_havoc_syntactic_general) + + + + + + + + + + + +subsection \Assignment rule\ + +subsubsection \Program expressions\ + +fun subst_pexp :: "var \ 'a pexp \ 'a pexp \ 'a pexp" where + "subst_pexp x e (PVar y) = (if x = y then e else PVar y)" +| "subst_pexp x e (PBinop p1 op p2) = PBinop (subst_pexp x e p1) op (subst_pexp x e p2)" +| "subst_pexp x e (PFun f p) = PFun f (subst_pexp x e p)" +| "subst_pexp _ _ e = e" (* Constants and quantified vars *) + +lemma subst_pexp_charact: + "interp_pexp (subst_pexp x e' e) \ = interp_pexp e (\(x := interp_pexp e' \))" +proof (induct e) + case (PVar x) + then show ?case + by (metis fun_upd_apply interp_pexp.simps(1) subst_pexp.simps(1)) +qed (simp_all) + +fun subst_pbexp :: "var \ 'a pexp \ 'a pbexp \ 'a pbexp" where + "subst_pbexp x e (PBAnd pb1 pb2) = PBAnd (subst_pbexp x e pb1) (subst_pbexp x e pb2)" +| "subst_pbexp x e (PBOr pb1 pb2) = PBOr (subst_pbexp x e pb1) (subst_pbexp x e pb2)" +| "subst_pbexp x e (PBComp p1 cmp p2) = PBComp (subst_pexp x e p1) cmp (subst_pexp x e p2)" +| "subst_pbexp _ _ (PBConst b) = PBConst b" + +lemma subst_pbexp_charact: + "interp_pbexp (subst_pbexp x e pb) \ \ interp_pbexp pb (\(x := interp_pexp e \))" +proof (induct pb) + case (PBComp x1 x2 x3) + then show ?case + using interp_pbexp.simps(4) subst_pexp_charact + by (metis subst_pbexp.simps(3)) +qed (simp_all) + + +subsubsection \Expressions (Boolean and values)\ + + +definition subst_all_states where + "subst_all_states x pe states = map (subst_state x pe) states" + +fun subst_exp :: "var \ 'a pexp \ 'a exp \ 'a exp" where + "subst_exp x pe (EPVar st y) = (if x = y then pexp_to_exp st pe else EPVar st y)" +| "subst_exp x pe (EBinop e1 bop e2) = EBinop (subst_exp x pe e1) bop (subst_exp x pe e2)" +| "subst_exp x pe (EFun f e) = EFun f (subst_exp x pe e)" +| "subst_exp _ _ e = e" (* Logical variables, quantified variables, constants *) + + +lemma subst_exp_charact_aux: + "snd (subst_state x pe (states ! st)) x = interp_exp vals states (pexp_to_exp st pe)" + by (induct pe) (simp_all add: subst_state_def) + +lemma subst_exp_charact: + assumes "wf_exp nv (length states) e" + shows "interp_exp vals states (subst_exp x pe e) = interp_exp vals (subst_all_states x pe states) e" + using assms +proof (induct e) + case (EPVar st y) + let ?states = "subst_all_states x pe states" + have "snd (subst_state x pe (states ! st)) = snd (?states ! st)" + by (metis EPVar.prems(1) nth_map subst_all_states_def wf_exp.simps(1)) + show "interp_exp vals states (subst_exp x pe (EPVar st y)) = interp_exp vals ?states (EPVar st y)" + proof (cases "x = y") + case True + then have "interp_exp vals states (subst_exp x pe (EPVar st y)) = interp_exp vals states (pexp_to_exp st pe)" + by simp + moreover have "interp_exp vals ?states (EPVar st y) = snd (?states ! st) y" + by simp + moreover have "... = snd (subst_state x pe (states ! st)) y" + by (simp add: \snd (subst_state x pe (states ! st)) = snd (subst_all_states x pe states ! st)\) + moreover have "snd (subst_state x pe (states ! st)) x = interp_exp vals states (pexp_to_exp st pe)" + by (metis subst_exp_charact_aux) + ultimately show ?thesis + using True by fastforce + next + case False + then show ?thesis + by (metis \snd (subst_state x pe (states ! st)) = snd (subst_all_states x pe states ! st)\ fun_upd_other interp_exp.simps(1) snd_conv subst_exp.simps(1) subst_state_def) + qed +next + case (ELVar st y) + let ?states = "subst_all_states x pe states" + have "fst (states ! st) = fst (?states ! st)" + by (metis ELVar.prems(1) fst_conv nth_map subst_all_states_def subst_state_def wf_exp.simps(2)) + have "interp_exp vals states (subst_exp x pe (ELVar st y)) = interp_exp vals states (ELVar st y)" + by simp + also have "... = fst (states ! st) y" + by simp + also have "... = fst (?states ! st) y" + by (simp add: \fst (states ! st) = fst (subst_all_states x pe states ! st)\) + also have "... = interp_exp vals ?states (ELVar st y)" + by auto + then show "interp_exp vals states (subst_exp x pe (ELVar st y)) = interp_exp vals ?states (ELVar st y)" + using calculation by presburger +qed (simp_all) + + + +subsubsection \Assertions\ + +fun transform_assign where + "transform_assign x pe (AForallState A) = AForallState (subst_assertion_single 0 x (pexp_to_exp 0 pe) (transform_assign x pe A))" +| "transform_assign x pe (AExistsState A) = AExistsState (subst_assertion_single 0 x (pexp_to_exp 0 pe) (transform_assign x pe A))" +| "transform_assign x pe (AExists A) = AExists (transform_assign x pe A)" +| "transform_assign x pe (AForall A) = AForall (transform_assign x pe A)" +| "transform_assign x pe (AOr A B) = AOr (transform_assign x pe A) (transform_assign x pe B)" +| "transform_assign x pe (AAnd A B) = AAnd (transform_assign x pe A) (transform_assign x pe B)" +| "transform_assign x pe (AConst b) = AConst b" +| "transform_assign x pe (AComp e1 cmp e2) = AComp e1 cmp e2" + + +lemma transform_assign_works: + "sat_assertion vals states (transform_assign x pe A) S = sat_assertion vals states A (subst_state x pe ` S)" +proof (induct A arbitrary: vals states) + case (AForallState A) + have "sat_assertion vals states (transform_assign x pe (AForallState A)) S + \ (\\\S. sat_assertion vals (\ # states) (subst_assertion_single 0 x (pexp_to_exp 0 pe) (transform_assign x pe A)) S)" + by auto + also have "... \ (\\\S. sat_assertion vals (update_state_at (\ # states) 0 x (interp_exp vals (\ # states) (pexp_to_exp 0 pe))) (transform_assign x pe A) S)" + by (simp add: subst_assertion_single_charact_better) + also have "... \ (\\\S. sat_assertion vals (update_state_at (\ # states) 0 x (interp_exp vals (\ # states) (pexp_to_exp 0 pe))) A (subst_state x pe ` S))" + using AForallState by presburger + also have "... \ (\\\S. sat_assertion vals (update_state \ x (interp_exp vals (\ # states) (pexp_to_exp 0 pe)) # states) A (subst_state x pe ` S))" + using update_state_at_def + by (metis list_update_code(2) nth_Cons_0) + also have "... \ (\\\S. sat_assertion vals (update_state \ x (interp_pexp pe (snd \)) # states) A (subst_state x pe ` S))" + by (metis nth_Cons_0 same_syn_sem_exp) + finally show "sat_assertion vals states (transform_assign x pe (AForallState A)) S = sat_assertion vals states (AForallState A) (subst_state x pe ` S)" + by (simp add: subst_state_def update_state_def) +next + case (AExistsState A) + have "sat_assertion vals states (transform_assign x pe (AExistsState A)) S + \ (\\\S. sat_assertion vals (\ # states) (subst_assertion_single 0 x (pexp_to_exp 0 pe) (transform_assign x pe A)) S)" + by auto + also have "... \ (\\\S. sat_assertion vals (update_state_at (\ # states) 0 x (interp_exp vals (\ # states) (pexp_to_exp 0 pe))) (transform_assign x pe A) S)" + by (simp add: subst_assertion_single_charact_better) + also have "... \ (\\\S. sat_assertion vals (update_state_at (\ # states) 0 x (interp_exp vals (\ # states) (pexp_to_exp 0 pe))) A (subst_state x pe ` S))" + using AExistsState by presburger + also have "... \ (\\\S. sat_assertion vals (update_state \ x (interp_exp vals (\ # states) (pexp_to_exp 0 pe)) # states) A (subst_state x pe ` S))" + using update_state_at_def + by (metis list_update_code(2) nth_Cons_0) + also have "... \ (\\\S. sat_assertion vals (update_state \ x (interp_pexp pe (snd \)) # states) A (subst_state x pe ` S))" + by (metis nth_Cons_0 same_syn_sem_exp) + finally show "sat_assertion vals states (transform_assign x pe (AExistsState A)) S = sat_assertion vals states (AExistsState A) (subst_state x pe ` S)" + by (simp add: subst_state_def update_state_def) +qed (simp_all) + + +subsubsection \Syntactic rule for assignments\ + +theorem rule_assign_syntactic_general: + "\ { sat_assertion vals states (transform_assign x pe P) } Assign x (interp_pexp pe) {sat_assertion vals states P}" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "sat_assertion vals states (transform_assign x pe P) S" + then have "sat_assertion vals states P (subst_state x pe ` S)" + using transform_assign_works by blast + moreover have "{(l, \(x := interp_pexp pe \)) |l \. (l, \) \ S} = subst_state x pe ` S" (is "?A = ?B") + proof + show "?A \ ?B" + proof + fix \ assume "\ \ ?A" + then obtain l \ where "\ = (l, \(x := interp_pexp pe \))" "(l, \) \ S" + by blast + then show "\ \ ?B" + by (metis (mono_tags, lifting) fst_conv image_iff snd_conv subst_state_def) + qed + show "?B \ ?A" + using subst_state_def by fastforce + qed + ultimately show "sat_assertion vals states P (sem (Assign x (interp_pexp pe)) S)" + by (simp add: sem_assign) +qed + + + +theorem rule_assign_syntactic: + "\ { interp_assert (transform_assign x pe P) } Assign x (interp_pexp pe) {interp_assert P}" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "interp_assert (transform_assign x pe P) S" + then have "sat_assertion [] [] P (subst_state x pe ` S)" + using transform_assign_works by blast + moreover have "{(l, \(x := interp_pexp pe \)) |l \. (l, \) \ S} = subst_state x pe ` S" (is "?A = ?B") + proof + show "?A \ ?B" + proof + fix \ assume "\ \ ?A" + then obtain l \ where "\ = (l, \(x := interp_pexp pe \))" "(l, \) \ S" + by blast + then show "\ \ ?B" + by (metis (mono_tags, lifting) fst_conv image_iff snd_conv subst_state_def) + qed + show "?B \ ?A" + using subst_state_def by fastforce + qed + ultimately show "interp_assert P (sem (Assign x (interp_pexp pe)) S)" + by (simp add: sem_assign) +qed + + + + + +subsection \Loop rules\ + + +fun no_exists_state :: "'a assertion \ bool" + where + "no_exists_state (AConst _) \ True" +| "no_exists_state (AComp _ _ _) \ True" +| "no_exists_state (AForallState A) \ no_exists_state A" +| "no_exists_state (AExistsState A) \ False" +| "no_exists_state (AForall A) \ no_exists_state A" +| "no_exists_state (AExists A) \ no_exists_state A" +| "no_exists_state (AAnd A B) \ no_exists_state A \ no_exists_state B" +| "no_exists_state (AOr A B) \ no_exists_state A \ no_exists_state B" + +lemma mono_sym_then_up_closed: + assumes "no_exists_state A" + and "S \ S'" + and "sat_assertion vals states A S'" + shows "sat_assertion vals states A S" + using assms +proof (induct A arbitrary: vals states) + case (AExistsState A) + then show ?case + by (metis no_exists_state.simps(4)) +qed (auto) + +definition up_closed where + "up_closed A \ (\S S' vals states. S \ S' \ sat_assertion vals states A S \ sat_assertion vals states A S')" + +lemma up_closedE: + assumes "up_closed A" + and "S \ S'" + and "sat_assertion vals states A S" + shows "sat_assertion vals states A S'" + using assms(1) assms(2) assms(3) up_closed_def by blast + + + +lemma sat_assertion_aforallstateI: + assumes "\\. \ \ S \ sat_assertion vals (\ # states) A S" + shows "sat_assertion vals states (AForallState A) S" + using assms sat_assertion.simps(3) by blast + + +lemma join_entails: + assumes "up_closed A" + and "sat_assertion vals states (AForallState A) S1" + and "sat_assertion vals states (AForallState A) S2" + shows "sat_assertion vals states (AForallState A) (S1 \ S2)" +proof (rule sat_assertion_aforallstateI) + fix \ + assume asm0: "\ \ S1 \ S2" + show "sat_assertion vals (\ # states) A (S1 \ S2)" + proof (cases "\ \ S1") + case True + then have "sat_assertion vals (\ # states) A S1" + using assms(2) sat_assertion.simps(3) by blast + then show ?thesis + by (meson assms(1) sup_ge1 up_closedE) + next + case False + then show ?thesis + by (metis UnE asm0 assms(1) assms(3) sat_assertion.simps(3) sup_ge2 up_closedE) + qed +qed + + +lemma general_join_entails: + assumes "up_closed A" + and "\x. sat_assertion vals states (AForallState A) (F x)" + shows "sat_assertion vals states (AForallState A) (\x. F x)" +proof (rule sat_assertion_aforallstateI) + fix \ + assume asm0: "\ \ \ (range F)" + then obtain x where "\ \ F x" + by blast + then have "sat_assertion vals (\ # states) A (F x)" + using assms(2) by force + then show "sat_assertion vals (\ # states) A (\ (range F))" + by (meson Union_upper assms(1) range_eqI up_closedE) +qed + + + +fun no_forall_state :: "'a assertion \ bool" + where + "no_forall_state (AConst _) \ True" +| "no_forall_state (AComp _ _ _) \ True" +| "no_forall_state (AForallState A) \ False" +| "no_forall_state (AExistsState A) \ no_forall_state A" +| "no_forall_state (AForall A) \ no_forall_state A" +| "no_forall_state (AExists A) \ no_forall_state A" +| "no_forall_state (AAnd A B) \ no_forall_state A \ no_forall_state B" +| "no_forall_state (AOr A B) \ no_forall_state A \ no_forall_state B" + +lemma no_forall_exists_state_not: + "no_forall_state A \ no_exists_state (ANot A)" + by (induct A) auto + + + +fun no_forall_state_after_existential :: "'a assertion \ bool" + where + "no_forall_state_after_existential (AConst _) \ True" +| "no_forall_state_after_existential (AComp _ _ _) \ True" +| "no_forall_state_after_existential (AForallState A) \ no_forall_state_after_existential A" +| "no_forall_state_after_existential (AForall A) \ no_forall_state_after_existential A" +| "no_forall_state_after_existential (AAnd A B) \ no_forall_state_after_existential A \ no_forall_state_after_existential B" +| "no_forall_state_after_existential (AOr A B) \ no_forall_state_after_existential A \ no_forall_state_after_existential B" +| "no_forall_state_after_existential (AExists A) \ no_forall_state A" +| "no_forall_state_after_existential (AExistsState A) \ no_forall_state A" + + +lemma up_closed_from_no_exists_state_false: + assumes "no_forall_state A" + and "sat_assertion vals states A (S n)" + shows "sat_assertion vals states A (\n. S n)" + using assms +proof (induct A) + case (AExistsState A) + then show ?case + by (meson UN_upper no_forall_exists_state_not iso_tuple_UNIV_I mono_sym_then_up_closed sat_assertion_Not) +next + case (AForall A) + then show ?case + by (meson UN_upper no_forall_exists_state_not iso_tuple_UNIV_I mono_sym_then_up_closed sat_assertion_Not) +next + case (AExists A) + then show ?case + by (meson UNIV_I no_forall_exists_state_not UN_upper mono_sym_then_up_closed sat_assertion_Not) +qed (force+) + + +definition shift_sequence :: "(nat \ 'a) \ nat \ (nat \ 'a)" + where + "shift_sequence S n m = S (m + n)" + +lemma shift_sequence_properties: + assumes "ascending S" + shows "ascending (shift_sequence S n)" + and "(\m. S m) = (\m. (shift_sequence S n) m)" (is "?A = ?B") + apply (metis (mono_tags, lifting) Suc_n_not_le_n add_Suc ascendingI ascending_def assms nle_le shift_sequence_def) +proof + show "?B \ ?A" + by (simp add: SUP_le_iff UN_upper shift_sequence_def) + show "?A \ ?B" + by (metis SUP_mono' ascending_def assms le_add1 shift_sequence_def) +qed + +fun extract_indices_sat_P where + "extract_indices_sat_P P S 0 = (SOME n. P (S n))" +| "extract_indices_sat_P P S (Suc m) = (SOME n. P (S n) \ n > extract_indices_sat_P P S m)" + +definition holds_infinitely_often where + "holds_infinitely_often P S \ (\m. \n. n > m \ P (S n))" + +lemma extract_indices_sat_P_properties: + assumes "holds_infinitely_often P S" + shows "P (S (extract_indices_sat_P P S 0))" + and "n > 0 \ P (S (extract_indices_sat_P P S n)) + \ extract_indices_sat_P P S n > extract_indices_sat_P P S (n - 1)" + apply (metis assms extract_indices_sat_P.simps(1) holds_infinitely_often_def someI_ex) + using assms +proof (induct n) + case (Suc n) + then have "P (S (extract_indices_sat_P P S n))" + by (metis bot_nat_0.not_eq_extremum extract_indices_sat_P.simps(1) holds_infinitely_often_def someI_ex) + let ?P = "\m. P (S m) \ m > extract_indices_sat_P P S n" + have "?P (SOME m. ?P m)" + proof (rule someI_ex) + show "\x. P (S x) \ extract_indices_sat_P P S n < x" + by (meson Suc.prems(2) holds_infinitely_often_def) + qed + then show ?case + by simp +qed (simp) + +lemma extract_indices_sat_P_larger: + assumes "holds_infinitely_often P S" + shows "extract_indices_sat_P P S n \ n" + using assms +proof (induct n) + case (Suc n) + then show ?case + by (metis Suc_leI diff_Suc_1 diff_diff_cancel extract_indices_sat_P_properties(2) less_imp_diff_less zero_less_Suc) +qed (simp) + +definition subseq_sat where + "subseq_sat P S n = S (extract_indices_sat_P P S n)" + +lemma subseq_sat_properties: + assumes "holds_infinitely_often P S" + and "ascending S" + shows "ascending (subseq_sat P S)" + and "\n. P (subseq_sat P S n)" + and "(\n. S n) = (\n. subseq_sat P S n)" (is "?A = ?B") +proof (rule ascendingI) + fix n + have "extract_indices_sat_P P S (Suc n) > extract_indices_sat_P P S n" + by (metis assms(1) diff_Suc_1 extract_indices_sat_P_properties(2) zero_less_Suc) + then show "subseq_sat P S n \ subseq_sat P S (Suc n)" + by (metis ascending_def assms(2) less_imp_le_nat subseq_sat_def) + fix n show "P (subseq_sat P S n)" + by (metis assms(1) extract_indices_sat_P_properties(1) extract_indices_sat_P_properties(2) gr0I subseq_sat_def) +next + show "?A = ?B" + proof + show "?B \ ?A" + by (simp add: SUP_le_iff UN_upper subseq_sat_def) + show "?A \ ?B" + proof + fix x assume "x \ \ (range S)" + then obtain n where "x \ S n" by blast + then have "x \ subseq_sat P S n" + by (metis ascending_def assms(1) assms(2) extract_indices_sat_P_larger subseq_sat_def subsetD) + then show "x \ ?B" by blast + qed + qed +qed + + +lemma no_forall_state_after_existential_sem: + assumes "no_forall_state_after_existential A" + and "ascending S" + and "\n. sat_assertion vals states A (S n)" + shows "sat_assertion vals states A (\n. S n)" + using assms +proof (induct A arbitrary: vals states S) + case (AForallState A) + show ?case + proof (rule sat_assertion_aforallstateI) + fix \ + assume "\ \ \ (range S)" + then obtain n where "\ \ S n" by blast + then have "\m. m \ n \ sat_assertion vals (\ # states) A (S m)" + by (meson AForallState.prems(2) AForallState.prems(3) ascending_def sat_assertion.simps(3) subsetD) + let ?S = "shift_sequence S n" + have "sat_assertion vals (\ # states) A (\ (range ?S))" + proof (rule AForallState(1)) + show "no_forall_state_after_existential A" + using AForallState.prems(1) by auto + show "ascending ?S" + by (simp add: AForallState.prems(2) shift_sequence_properties(1)) + fix m show "sat_assertion vals (\ # states) A (shift_sequence S n m)" + by (simp add: \\m. n \ m \ sat_assertion vals (\ # states) A (S m)\ shift_sequence_def) + qed + then show "sat_assertion vals (\ # states) A (\ (range S))" + by (metis AForallState.prems(2) shift_sequence_properties(2)) + qed +next + case (AExistsState A) + then show ?case + by (meson no_forall_state.simps(4) no_forall_state_after_existential.simps(8) up_closed_from_no_exists_state_false) +next + case (AExists A) + then show ?case + by (meson no_forall_state.simps(6) no_forall_state_after_existential.simps(7) up_closed_from_no_exists_state_false) +next + case (AOr A1 A2) +(* either A1 is infinitely often true, or either A2 is... *) + show ?case + proof (cases "holds_infinitely_often (sat_assertion vals states A1) S") + case True + then have "sat_assertion vals states A1 (\ (range (subseq_sat (sat_assertion vals states A1) S)))" + using AOr.hyps(1) AOr.prems(1) AOr.prems(2) subseq_sat_properties(1) subseq_sat_properties(2) no_forall_state_after_existential.simps(6) by blast + then show ?thesis + using AOr.prems(2) True subseq_sat_properties(3) by fastforce + next + case False + then have "holds_infinitely_often (sat_assertion vals states A2) S" + using AOr.prems(3) sat_assertion.simps(8) holds_infinitely_often_def + by (metis gt_ex max_less_iff_conj) + then have "sat_assertion vals states A2 (\ (range (subseq_sat (sat_assertion vals states A2) S)))" + using AOr.hyps(2) AOr.prems(1) AOr.prems(2) subseq_sat_properties(1) subseq_sat_properties(2) no_forall_state_after_existential.simps(6) by blast + then show ?thesis + using AOr.prems(2) \holds_infinitely_often (sat_assertion vals states A2) S\ subseq_sat_properties(3) by fastforce + qed +next + case (AAnd A1 A2) + then show ?case + using sat_assertion.simps(7) no_forall_state_after_existential.simps(5) by blast +qed (simp_all) + + +(* Can be used with while rule! *) +lemma upwards_closed_syn_sem_practical: + assumes "no_forall_state_after_existential A" + shows "upwards_closed (\n. interp_assert A) (interp_assert A)" + by (simp add: assms no_forall_state_after_existential_sem upwards_closed_def) + +theorem while_general_syntactic: + assumes "\n. \ {P n} if_then b C {P (Suc n)}" + and "\n. \ {P n} Assume (lnot b) {interp_assert A}" + and "no_forall_state_after_existential A" + shows "\ {P 0} while_cond b C {conj (interp_assert A) (holds_forall (lnot b))}" + by (metis assms(1) assms(2) assms(3) upwards_closed_syn_sem_practical while_general) + +theorem while_forall_exists_simpler: + assumes "\ {I} if_then b C {I}" + and "\ {I} Assume (lnot b) {interp_assert Q}" + and "no_forall_state_after_existential Q" + shows "\ {I} while_cond b C {conj (interp_assert Q) (holds_forall (lnot b))}" + by (metis assms(1) assms(2) assms(3) upwards_closed_syn_sem_practical while_general) + +(* Side-conditions: No normal existential on the left of a forallstate, no existential state *) +theorem while_d_syntactic: + assumes "\ { interp_assert A } if_then b C { interp_assert A }" + and "no_forall_state_after_existential A" + and "no_exists_state A" + shows "\ { interp_assert A } while_cond b C {conj (interp_assert A) (holds_forall (lnot b))}" + using assms(1) +proof (rule while_d) + show "upwards_closed (\a. interp_assert A) (interp_assert A)" + by (simp add: assms(2) no_forall_state_after_existential_sem upwards_closed_def) + fix n show "downwards_closed (interp_assert A)" + using assms(3) downwards_closed_def mono_sym_then_up_closed by fastforce +qed + + + +lemma downwards_closed_is_hypersafety: + "hypersafety P \ downwards_closed P" + using downwards_closed_def hypersafety_def by metis + + +subsection \Rewrite rules for 'a assertions\ + +definition equiv where + "equiv A B \ (\vals states S. sat_assertion vals states A S \ sat_assertion vals states B S)" + +lemma forall_commute: + "equiv (AForallState (AForall A)) (AForall (AForallState A))" + using equiv_def by force + +lemma exists_commute: + "equiv (AExistsState (AExists A)) (AExists (AExistsState A))" + using equiv_def by force + +lemma forall_state_and: + "equiv (AForallState (AAnd A B)) (AAnd (AForallState A) (AForallState B))" + using equiv_def by force + +lemma exists_state_or: + "equiv (AExistsState (AOr A B)) (AOr (AExistsState A) (AExistsState B))" + using equiv_def by force + +lemma forall_and: + "equiv (AForall (AAnd A B)) (AAnd (AForall A) (AForall B))" + using equiv_def by force + +lemma exists_or: + "equiv (AExists (AOr A B)) (AOr (AExists A) (AExists B))" + using equiv_def by force + +lemma entailment_natural_partition: + assumes "no_forall_state P" + shows "entails (natural_partition (\(n::nat). interp_assert (AForallState P))) (interp_assert (AForallState P))" +proof (rule entailsI) + fix S :: "((nat \ 'a) \ (nat \ 'a)) set" + + assume asm0: "natural_partition (\(n::nat). interp_assert (AForallState P)) S" + then obtain F where "S = (\(n::nat). F n)" "\n. interp_assert (AForallState P) (F n)" + using natural_partitionE[of "\n. interp_assert (AForallState P)" S] by blast + then have "\\. \ \ S \ sat_assertion [] [\] P S" + by (meson UN_iff assms sat_assertion.simps(3) up_closed_from_no_exists_state_false) + then show "interp_assert (AForallState P) S" + by simp +qed + + +lemma no_forall_state_mono: + assumes "no_forall_state A" + and "sat_assertion vals states A S" + and "S \ S'" + shows "sat_assertion vals states A S'" +by (metis assms(1) assms(2) assms(3) mono_sym_then_up_closed no_forall_exists_state_not sat_assertion_Not) + + +lemma entailment_loop_join: + assumes "no_forall_state P" + shows "entails (join (interp_assert (AForallState P)) (interp_assert (AForallState P))) (interp_assert (AForallState P))" +proof (rule entailsI) + fix S :: "((nat \ 'a) \ (nat \ 'a)) set" + + assume asm0: "join (interp_assert (AForallState P)) (interp_assert (AForallState P)) S" + then obtain S1 S2 where r: "S = S1 \ S2" "interp_assert (AForallState P) S1" "interp_assert (AForallState P) S2" + by (metis join_def) + have "\\. \ \ S \ sat_assertion [] [\] P S" + proof - + fix \ assume asm1: "\ \ S" + then show "sat_assertion [] [\] P S" + by (metis Un_iff r assms no_forall_state_mono sat_assertion.simps(3) sup_ge1 sup_ge2) + qed + then show "interp_assert (AForallState P) S" + by simp +qed + + + + +subsection \Free variables and safe frame rule\ + +fun wr :: "(nat, nat) stmt \ nat set" where + "wr Skip = {}" +| "wr (Assign x _) = {x}" +| "wr (Havoc x) = {x}" +| "wr (Assume b) = {}" +| "wr (C1;; C2) = wr C1 \ wr C2" +| "wr (If C1 C2) = wr C1 \ wr C2" +| "wr (While C) = wr C" + + +definition agree_on where + "agree_on V \ \' \ (\x\V. \ x = \' x)" + +lemma agree_onI: + assumes "\x. x\V \ \ x = \' x" + shows "agree_on V \ \'" + using agree_on_def assms by blast + +lemma agree_onE: + assumes "agree_on V \ \'" + and "x \ V" + shows "\ x = \' x" + by (meson agree_on_def assms(1) assms(2)) + +lemma agree_on_subset: + assumes "agree_on V' \ \'" + and "V \ V'" + shows "agree_on V \ \'" + by (meson agree_on_def assms(1) assms(2) in_mono) + +lemma agree_on_trans: + assumes "agree_on V \ \'" + and "agree_on V \' \''" + shows "agree_on V \ \''" +proof (rule agree_onI) + fix x assume "x \ V" + then show "\ x = \'' x" + by (metis agree_on_def assms(1) assms(2)) +qed + +lemma agree_on_sym: + assumes "agree_on V \ \'" + shows "agree_on V \' \" + by (metis agree_on_def assms) + +lemma wr_charact: + assumes "single_sem C \ \'" + and "wr C \ V = {}" + shows "agree_on V \ \'" + using assms +proof (induct rule: single_sem.induct) + case (SemSeq C1 \ \1 C2 \2) + then show ?case + by (metis Un_empty agree_on_trans inf_sup_distrib2 wr.simps(5)) +qed (auto simp add: agree_on_def) + +fun fv_exp :: "'a exp \ var set" where + "fv_exp (EBinop e1 _ e2) = fv_exp e1 \ fv_exp e2" +| "fv_exp (EPVar _ x) = {x}" +| "fv_exp (EFun _ e) = fv_exp e" +| "fv_exp _ = {}" + +lemma fv_wr_charact_exp: +(* all program variables *) + assumes "agree_on (fv_exp e) \ \'" + and "n \ length states" + and "wf_exp nv (Suc (length states)) e" + shows "interp_exp vals (insert_at n (l, \) states) e = interp_exp vals (insert_at n (l, \') states) e" + using assms +proof (induct e) + case (EPVar st x) + then show ?case + proof (cases "st < n") + case True + then show ?thesis + by (simp add: assms(2) insert_at_charact_1) + next + case False + then have "st \ n" by simp + then show ?thesis + proof (cases "st = n") + case True + then have "interp_exp vals (insert_at n (l, \) states) (EPVar st x) = \ x" + by (simp add: assms(2) insert_at_charact_2) + also have "... = \' x" + by (metis EPVar.prems(1) agree_on_def fv_exp.simps(2) insertCI) + then show ?thesis + by (simp add: True assms(2) insert_at_charact_2) + next + case False + then have "interp_exp vals (insert_at n (l, \) states) (EPVar st x) = snd ((insert_at n (l, \) states) ! st) x" + by simp + also have "... = snd (states ! (st - 1)) x" + by (metis False One_nat_def \n \ st\ add_diff_inverse_nat assms(2) dual_order.antisym insert_at_charact_3 le_less_Suc_eq le_zero_eq not_less_eq_eq plus_1_eq_Suc) + then show ?thesis + by (metis False One_nat_def \n \ st\ add_diff_inverse_nat assms(2) dual_order.antisym insert_at_charact_3 interp_exp.simps(1) le_less_Suc_eq le_zero_eq not_less_eq_eq plus_1_eq_Suc) + qed + qed +next + case (ELVar st x) + then show ?case + proof (cases "st < n") + case True + then show ?thesis + by (simp add: assms(2) insert_at_charact_1) + next + case False + then have "st \ n" by simp + then show ?thesis + proof (cases "st = n") + case True + then have "interp_exp vals (insert_at n (l, \) states) (ELVar st x) = l x" + by (simp add: assms(2) insert_at_charact_2) + then show ?thesis + by (simp add: True assms(2) insert_at_charact_2) + next + case False + then have "interp_exp vals (insert_at n (l, \) states) (ELVar st x) = fst ((insert_at n (l, \) states) ! st) x" + by simp + also have "... = fst (states ! (st - 1)) x" + by (metis False One_nat_def \n \ st\ add_diff_inverse_nat assms(2) dual_order.antisym insert_at_charact_3 le_less_Suc_eq le_zero_eq not_less_eq_eq plus_1_eq_Suc) + then show ?thesis + by (metis False One_nat_def \n \ st\ add_diff_inverse_nat assms(2) dual_order.antisym insert_at_charact_3 interp_exp.simps(2) le_less_Suc_eq le_zero_eq not_less_eq_eq plus_1_eq_Suc) + qed + qed +next + case (EBinop e1 x2 e2) + then show ?case + by (simp add: agree_on_def) +qed (simp_all) + + +(* program variables... *) +fun fv where + "fv (AAnd F1 F2) = fv F1 \ fv F2" +| "fv (AOr F1 F2) = fv F1 \ fv F2" +| "fv (AForall F) = fv F" +| "fv (AExists F) = fv F" +| "fv (AForallState F) = fv F" +| "fv (AExistsState F) = fv F" +| "fv (AConst b) = {}" +| "fv (AComp e1 cmp e2) = fv_exp e1 \ fv_exp e2" + + + + +lemma fv_wr_charact_aux: +(* all program variables *) + assumes "agree_on (fv F) \ \'" + and "n \ length states" + and "sat_assertion vals (insert_at n (l, \) states) F S" + and "wf_assertion_aux nv (Suc (length states)) F" + shows "sat_assertion vals (insert_at n (l, \') states) F S" + using assms +proof (induct F arbitrary: vals states n nv) + case (AExists F) + then show ?case + by (metis fv.simps(4) sat_assertion.simps(6) wf_assertion_aux.simps(6)) +next + case (AComp e1 cmp e2) + then have "interp_exp vals (insert_at n (l, \) states) e1 = interp_exp vals (insert_at n (l, \') states) e1" + using fv_wr_charact_exp[of e1 \ \' n states nv vals l] agree_on_subset by fastforce + moreover have "interp_exp vals (insert_at n (l, \) states) e2 = interp_exp vals (insert_at n (l, \') states) e2" + using fv_wr_charact_exp[of e2 \ \' n states nv vals l] agree_on_subset AComp by auto + ultimately show "sat_assertion vals (insert_at n (l, \') states) (AComp e1 cmp e2) S" + using AComp.prems(3) by auto +next + case (AForallState F) + then have "\\. \ \ S \ sat_assertion vals (insert_at (Suc n) (l, \) (\ # states)) F S" + by simp + then have "\\. \ \ S \ sat_assertion vals (insert_at (Suc n) (l, \') (\ # states)) F S" + by (metis AForallState.hyps AForallState.prems(1) AForallState.prems(2) AForallState.prems(4) Suc_le_mono fv.simps(5) length_Cons wf_assertion_aux.simps(7)) + then show ?case by simp +next + case (AExistsState F) + then obtain \ where asm0: "\ \ S" "sat_assertion vals (insert_at (Suc n) (l, \) (\ # states)) F S" + by auto + then have "sat_assertion vals (insert_at (Suc n) (l, \') (\ # states)) F S" + by (metis AExistsState.hyps AExistsState.prems(1) AExistsState.prems(2) AExistsState.prems(4) Suc_le_mono fv.simps(6) length_Cons wf_assertion_aux.simps(8)) + then show ?case using asm0 by auto +qed (auto simp add: agree_on_def) + + +lemma fv_wr_charact: + assumes "agree_on (fv F) \ \'" + and "sat_assertion vals ((l, \) # states) F S" + and "wf_assertion_aux nv (Suc (length states)) F" + shows "sat_assertion vals ((l, \') # states) F S" +proof - + have "sat_assertion vals (insert_at 0 (l, \') states) F S" + using assms fv_wr_charact_aux by fastforce + then show ?thesis + by simp +qed + + +lemma syntactic_safe_frame_preserved: + assumes "wr C \ fv F = {}" + and "sat_assertion vals states F S" + and "wf_assertion_aux nv (length states) F" + and "no_exists_state F" + shows "sat_assertion vals states F (sem C S)" + using assms +proof (induct F arbitrary: vals states nv) + case (AForallState F) + then have "\\. \ \ sem C S \ sat_assertion vals (\ # states) F (sem C S)" + proof - + fix \ assume asm0: "\ \ sem C S" + then obtain \ where "single_sem C \ (snd \)" "(fst \, \) \ S" + using in_sem by blast + then have "sat_assertion vals ((fst \, \) # states) F (sem C S)" + using AForallState.hyps AForallState.prems assms(1) by auto + moreover have "agree_on (fv F) \ (snd \)" + using AForallState.prems(1) \\(C::(nat, nat) stmt), \::nat \ nat\ \ snd (\::(nat \ nat) \ (nat \ nat))\ wr_charact by auto + moreover have "wf_assertion_aux nv (Suc (length states)) F" + using AForallState.prems(3) by auto + ultimately have "sat_assertion vals ((fst \, snd \) # states) F (sem C S)" + using fv_wr_charact[of F \ "snd \" vals "fst \" states "sem C S"] + by fast + then show "sat_assertion vals (\ # states) F (sem C S)" + by simp + qed + then show ?case + using AForallState.hyps AForallState.prems(2) assms(1) by auto +qed (fastforce+) + +theorem safe_frame_rule_syntactic: + assumes "wr C \ fv F = {}" + and "wf_assertion F" + and "no_exists_state F" + shows "\ {interp_assert F} C {interp_assert F}" + by (metis assms(1) assms(2) assms(3) hyper_hoare_tripleI list.size(3) syntactic_safe_frame_preserved) + +theorem LUpdateS: + assumes "\ { (\S. P S \ e_recorded_in_t e t S) } C { Q }" + and "not_fv_hyper t P" + and "not_fv_hyper t Q" + shows "\ { P } C { Q }" +proof (rule hyper_hoare_tripleI) + fix S assume asm: "P S" + let ?S = "assign_exp_to_lvar_set e t S" + have "Q (sem C ?S)" + using asm assms(1) assms(2) e_recorded_in_t_if_assigned hyper_hoare_tripleE not_fv_hyper_assign_exp by fastforce + then show "Q (sem C S)" + by (meson assign_exp_to_lvar_set_same_mod_updates assms(3) not_fv_hyper_def same_mod_updates_sym sem_update_commute) +qed + + +end \ No newline at end of file diff --git a/thys/HyperHoareLogic/TotalLogic.thy b/thys/HyperHoareLogic/TotalLogic.thy new file mode 100644 --- /dev/null +++ b/thys/HyperHoareLogic/TotalLogic.thy @@ -0,0 +1,1105 @@ +theory TotalLogic + imports Loops Compositionality SyntacticAssertions +begin + +section \Terminating Hyper-Triples\ + +definition total_hyper_triple ("\TERM {_} _ {_}" [51,0,0] 81) where + "\TERM {P} C {Q} \ ( \ {P} C {Q} \ (\S. P S \ (\\ \ S. \\'. single_sem C (snd \) \' )))" + +lemma total_hyper_triple_equiv: + "\TERM {P} C {Q} \ ( \ {P} C {Q} \ (\S. P S \ (\\ \ S. \\'. (fst \, \') \ sem C S \ single_sem C (snd \) \' )))" + by (metis prod.collapse single_step_then_in_sem total_hyper_triple_def) + +lemma total_hyper_tripleI: + assumes "\ {P} C {Q}" + and "\\ S. P S \ \ \ S \ (\\'. single_sem C (snd \) \' )" + shows "\TERM {P} C {Q}" + by (simp add: assms(1) assms(2) total_hyper_triple_def) + +definition terminates_in where + "terminates_in C S \ (\\ \ S. \\'. single_sem C (snd \) \')" + +lemma terminates_inI: + assumes "\\. \ \ S \ \\'. single_sem C (snd \) \'" + shows "terminates_in C S" + using assms terminates_in_def by blast + +lemma iterate_sem_mono: + assumes "S \ S'" + shows "iterate_sem n C S \ iterate_sem n C S'" + using assms + by (induct n arbitrary: S S') (simp_all add: sem_monotonic) + +lemma terminates_in_while_loop: + assumes "wfP lt" + and "\\ n. \ \ iterate_sem n (Assume b;; C) S \ b (snd \) \ (\\'. single_sem C (snd \) \' \ (\ b \' \ lt (e \') (e (snd \))))" + shows "terminates_in (while_cond b C) S" +proof (rule terminates_inI) + fix \ assume asm0: "\ \ S" + let ?S = "{\}" + let ?R = "{(x, y). lt x y}" + let ?Q = "{ e (snd \) |\ n. b (snd \) \ \ \ iterate_sem n (Assume b;; C) ?S}" + + show "\\'. \while_cond b C, snd \\ \ \'" + proof (cases "b (snd \)") + case False + then show ?thesis + by (metis SemAssume SemSeq SemWhileExit lnot_def while_cond_def) + next + case True + show ?thesis + proof (rule wfE_min) + show "wf ?R" + using assms(1) wfP_def by blast + show "e (snd \) \ ?Q" + using True asm0 iterate_sem.simps(1) by fastforce + fix z assume asm1: "z \ ?Q" "(\y. (y, z) \ {(x, y). lt x y} \ y \ ?Q)" + then obtain \' n where "z = e (snd \')" "b (snd \')" "\' \ iterate_sem n (Assume b;; C) ?S" + by blast + then obtain \' where "single_sem C (snd \') \' \ (\ b \' \ lt (e \') (e (snd \')))" + using assms(2) iterate_sem_mono[of ?S S n "Assume b;; C"] + by (meson asm0 empty_subsetI in_mono insert_subsetI) + then have "\ b \' \ e \' \ ?Q" + using \z = e (snd \')\ asm1(2) by blast + moreover have "(fst \', \') \ sem (Assume b;; C) (iterate_sem n (Assume b;; C) ?S)" + by (metis (no_types, opaque_lifting) SemAssume SemSeq \(\C, snd \'\ \ \') \ (\ b \' \ lt (e \') (e (snd \')))\ \\' \ iterate_sem n (Assume b ;; C) ?S\ \b (snd \')\ single_step_then_in_sem surjective_pairing) + ultimately have "\ b \'" + using iterate_sem.simps(2)[of n "Assume b;; C" "{\}"] mem_Collect_eq snd_conv + by (metis (mono_tags, lifting)) + then have "(fst \', \') \ iterate_sem (Suc n) (Assume b;; C) ?S" + by (simp add: \(fst \', \') \ sem (Assume b ;; C) (iterate_sem n (Assume b ;; C) ?S)\) + then have "(fst \', \') \ (\n. iterate_sem n (Assume b;; C) ?S)" by blast + then have "(fst \', \') \ filter_exp (lnot b) (\n. iterate_sem n (Assume b;; C) ?S)" + using filter_exp_def[of "lnot b"] lnot_def[of b \'] \\ b \'\ by force + then have "(fst \', \') \ sem (while_cond b C) ?S" + by (simp add: assume_sem filter_exp_def sem_seq sem_while while_cond_def) + then show "\\'. \while_cond b C, snd \\ \ \'" + by (metis in_sem singletonD snd_conv) + qed + qed +qed + +lemma total_hyper_triple_altI: + assumes "\S. P S \ Q (sem C S)" + and "\S. P S \ terminates_in C S" + shows "\TERM {P} C {Q}" + by (metis assms(1) assms(2) hyper_hoare_tripleI terminates_in_def total_hyper_triple_def) + + + +lemma syntactic_frame_preserved: + assumes "terminates_in C S" + and "wr C \ fv F = {}" + and "sat_assertion vals states F S" + and "wf_assertion_aux nv (length states) F" + shows "sat_assertion vals states F (sem C S)" + using assms +proof (induct F arbitrary: vals states nv) + case (AForallState F) + then have "\\. \ \ sem C S \ sat_assertion vals (\ # states) F (sem C S)" + proof - + fix \ assume asm0: "\ \ sem C S" + then obtain \ where "single_sem C \ (snd \)" "(fst \, \) \ S" + using in_sem by blast + then have "sat_assertion vals ((fst \, \) # states) F (sem C S)" + using AForallState.hyps AForallState.prems assms(1) by auto + moreover have "agree_on (fv F) \ (snd \)" + using AForallState.prems(2) \\(C::(nat, nat) stmt), \::nat \ nat\ \ snd (\::(nat \ nat) \ (nat \ nat))\ wr_charact by auto + moreover have "wf_assertion_aux nv (Suc (length states)) F" + using AForallState.prems(4) by auto + ultimately have "sat_assertion vals ((fst \, snd \) # states) F (sem C S)" + using fv_wr_charact[of F \ "snd \" vals "fst \" states "sem C S"] + by fast + then show "sat_assertion vals (\ # states) F (sem C S)" + by simp + qed + then show ?case + using AForallState.hyps AForallState.prems(2) assms(1) by auto +next + case (AExistsState F) + then obtain \ where asm0: "\ \ S" "sat_assertion vals (\ # states) F S" + by auto + then obtain \' where "single_sem C (snd \) \'" + using assms(1) terminates_in_def by blast + then have "sat_assertion vals ((fst \, \') # states) F S" + by (metis AExistsState.prems(2) AExistsState.prems(4) asm0(2) fv.simps(6) fv_wr_charact prod.collapse wf_assertion_aux.simps(8) wr_charact) + then show ?case + by (metis AExistsState.hyps AExistsState.prems(2) AExistsState.prems(4) \\C, snd \\ \ \'\ asm0(1) assms(1) fv.simps(6) length_Cons prod.collapse sat_assertion.simps(4) single_step_then_in_sem wf_assertion_aux.simps(8)) +qed (fastforce+) + + +theorem frame_rule_syntactic: + assumes "\TERM {P} C {Q}" + and "wr C \ fv F = {}" (* free *program* variables *) + and "wf_assertion F" (* No unbound free variable *) + shows "\TERM {conj P (interp_assert F)} C {conj Q (interp_assert F)}" +proof (rule total_hyper_tripleI) + let ?F = "interp_assert F" + show "\\ S. Logic.conj P ?F S \ \ \ S \ \\'. \C, snd \\ \ \'" + by (metis assms(1) conj_def total_hyper_triple_def) + show "\ {Logic.conj P ?F} C {Logic.conj Q ?F}" + proof (rule hyper_hoare_tripleI) + fix S assume asm0: "Logic.conj P ?F S" + then have "terminates_in C S" + by (simp add: \\\ S. Logic.conj P ?F S \ \ \ S \ \\'. \C, snd \\ \ \'\ terminates_in_def) + moreover have "?F (sem C S)" + by (metis asm0 assms(2) assms(3) calculation conj_def list.size(3) syntactic_frame_preserved) + ultimately show "Logic.conj Q ?F (sem C S)" + by (metis asm0 assms(1) conj_def hyper_hoare_tripleE total_hyper_triple_def) + qed +qed + + + + +subsection \Specialize rule\ + + +definition same_syn_sem_all :: "'a assertion \ ((nat, 'a, nat, 'a) state \ bool) \ bool" + where + "same_syn_sem_all bsyn bsem \ + (\states vals S. length states > 0 \ bsem (hd states) = sat_assertion vals states bsyn S)" + +lemma same_syn_sem_allI: + assumes "\states vals S. length states > 0 \ bsem (hd states) \ sat_assertion vals states bsyn S" + shows "same_syn_sem_all bsyn bsem" + by (simp add: assms same_syn_sem_all_def) + +lemma transform_assume_valid: + assumes "same_syn_sem_all bsyn bsem" + shows "sat_assertion vals states A (Set.filter bsem S) + \ sat_assertion vals states (transform_assume bsyn A) S" +proof (induct A arbitrary: vals states) + case (AForallState A) + let ?S = "Set.filter (bsem) S" + let ?A = "transform_assume bsyn A" + have "sat_assertion vals states (AForallState A) ?S \ (\\\?S. sat_assertion vals (\ # states) A ?S)" + by force + also have "... \ (\\\?S. sat_assertion vals (\ # states) ?A S)" + using AForallState by presburger + also have "... \ (\\\S. bsem \ \ sat_assertion vals (\ # states) ?A S)" + by fastforce + also have "... \ (\\\S. sat_assertion vals (\ # states) bsyn S \ sat_assertion vals (\ # states) ?A S)" + using assms same_syn_sem_all_def[of bsyn bsem] by auto + also have "... \ (\\\S. sat_assertion vals (\ # states) (AImp bsyn ?A) S)" + using sat_assertion_Imp by fast + also have "... \ sat_assertion vals states (AForallState (AImp bsyn ?A)) S" + using sat_assertion.simps(2) by force + then show ?case + using calculation transform_assume.simps(1) by fastforce +next + case (AExistsState A) + let ?S = "Set.filter (bsem) S" + let ?A = "transform_assume bsyn A" + have "sat_assertion vals states (AExistsState A) ?S \ (\\\?S. sat_assertion vals (\ # states) A ?S)" + by force + also have "... \ (\\\?S. sat_assertion vals (\ # states) ?A S)" + using AExistsState by presburger + also have "... \ (\\\S. bsem \ \ sat_assertion vals (\ # states) ?A S)" + by force + also have "... \ (\\\S. sat_assertion vals (\ # states) bsyn S \ sat_assertion vals (\ # states) ?A S)" + using assms same_syn_sem_all_def[of bsyn bsem] by auto + also have "... \ (\\\S. sat_assertion vals (\ # states) (AAnd bsyn ?A) S)" + by simp + also have "... \ sat_assertion vals states (AExistsState (AAnd bsyn ?A)) S" + using sat_assertion.simps(3) by force + then show ?case + using calculation transform_assume.simps(2) by fastforce +next + case (AForall A) + let ?S = "Set.filter (bsem) S" + let ?A = "transform_assume bsyn A" + have "sat_assertion vals states (AForall A) ?S \ (\v. sat_assertion (v # vals) states A ?S)" + by force + also have "... \ (\v. sat_assertion (v # vals) states ?A S)" + using AForall by presburger + also have "... \ sat_assertion vals states (AForall ?A) S" + by simp + then show ?case + using calculation transform_assume.simps(3) by fastforce +next + case (AExists A) + let ?S = "Set.filter (bsem) S" + let ?A = "transform_assume bsyn A" + have "sat_assertion vals states (AExists A) ?S \ (\v. sat_assertion (v # vals) states A ?S)" + by force + also have "... \ (\v. sat_assertion (v # vals) states ?A S)" + using AExists by presburger + also have "... \ sat_assertion vals states (AExists ?A) S" + by simp + then show ?case + using calculation transform_assume.simps(4) by fastforce +qed (simp_all) + + +fun indep_of_set where + "indep_of_set (AForall A) \ indep_of_set A" +| "indep_of_set (AExists A) \ indep_of_set A" +| "indep_of_set (AOr A B) \ indep_of_set A \ indep_of_set B" +| "indep_of_set (AAnd A B) \ indep_of_set A \ indep_of_set B" +| "indep_of_set (AComp _ _ _) \ True" +| "indep_of_set (AConst _) \ True" +| "indep_of_set (AForallState _) \ False" +| "indep_of_set (AExistsState _) \ False" + +lemma indep_of_set_charact: + assumes "indep_of_set A" + and "sat_assertion vals states A S" + shows "sat_assertion vals states A S'" + using assms +by (induct A arbitrary: vals states) (auto) + +lemma wf_exp_take: + assumes "wf_exp nv ns e" + shows "interp_exp vals states e = interp_exp (take nv vals) (take ns states) e" + using assms +proof (induct e arbitrary: nv ns vals states) + case (EQVar x) + then show ?case + by force +next + case (EBinop e1 x2 e2) + then show ?case + by (metis interp_exp.simps(5) wf_exp.simps(4)) +next + case (EFun f e) + then show ?case + by (metis interp_exp.simps(6) wf_exp.simps(5)) +qed (simp_all) + +lemma wf_assertion_aux_take: + assumes "wf_assertion_aux nv ns A" + shows "sat_assertion vals states A S \ sat_assertion (take nv vals) (take ns states) A S" + using assms +proof (induct A arbitrary: nv ns vals states) + case (AConst x) + then show ?case + by simp +next + case (AComp x1a x2 x3a) + then show ?case + by (metis sat_assertion.simps(2) wf_assertion_aux.simps(2) wf_exp_take) +next + case (AForallState A) + then show ?case using AForallState.hyps[of nv "Suc ns" vals] + by (metis sat_assertion.simps(3) take_Suc_Cons wf_assertion_aux.simps(7)) +next + case (AExistsState A) + then show ?case using AExistsState.hyps[of nv "Suc ns" vals] + by (metis sat_assertion.simps(4) take_Suc_Cons wf_assertion_aux.simps(8)) +next + case (AForall A) + then show ?case using AForall.hyps[of "Suc nv" ns _ states] + by (metis sat_assertion.simps(5) take_Suc_Cons wf_assertion_aux.simps(5)) +next + case (AExists A) + then show ?case using AExists.hyps[of "Suc nv" ns _ states] + by (metis sat_assertion.simps(6) take_Suc_Cons wf_assertion_aux.simps(6)) +next + case (AOr A1 A2) + then show ?case + by (metis sat_assertion.simps(8) wf_assertion_aux.simps(4)) +next + case (AAnd A1 A2) + then show ?case + by (metis sat_assertion.simps(7) wf_assertion_aux.simps(3)) +qed + +lemma syntactic_charact_for_equivalence: + assumes "indep_of_set A" + and "wf_assertion_aux (0::nat) (1::nat) A" + shows "sat_assertion vals (\ # states) A S \ sat_assertion [] [\] A {}" (is "?A \ ?B") +proof - + have "?A \ sat_assertion vals (\ # states) A {}" + using assms(1) indep_of_set_charact by blast + also have "... \ sat_assertion (take (0::nat) vals) (take (1::nat) (\ # states)) A {}" + using wf_assertion_aux_take[of 0 1 A vals "\ # states" "{}"] assms(2) + by blast + ultimately show ?thesis by simp +qed + + +definition get_bsem where + "get_bsem bsyn \ \ sat_assertion [] [\] bsyn {}" + +lemma syntactic_charact_for_bsem: + assumes "indep_of_set A" + and "wf_assertion_aux (0::nat) (1::nat) A" + shows "same_syn_sem_all A (get_bsem A)" +proof (rule same_syn_sem_allI) + fix states :: "((nat \ 'a) \ (nat \ 'a)) list" + fix vals S + assume asm0: "0 < length states" + then show "get_bsem A (hd states) = sat_assertion vals states A S" + by (metis assms(1) assms(2) get_bsem_def length_greater_0_conv list.collapse syntactic_charact_for_equivalence) +qed + + + +lemma get_bsem_is_bsem: + assumes "same_syn_sem_all bsyn bsem" + shows "bsem = get_bsem bsyn" +proof (rule ext) + fix x + have "bsem (hd [x]) = sat_assertion [] [x] bsyn {}" + using assms same_syn_sem_all_def by fastforce + then show "bsem x = get_bsem bsyn x" + by (simp add: get_bsem_def) +qed + + +lemma free_vars_syn_sem: + assumes "same_syn_sem_all bsyn bsem" + and "fst \ = fst \'" + and "agree_on (fv bsyn) (snd \) (snd \')" + and "bsem \" + and "wf_assertion_aux 0 (Suc 0) bsyn" + shows "bsem \'" +proof - + have "sat_assertion [] (insert_at 0 (fst \, snd \') []) bsyn {}" + using assms(3) + proof (rule fv_wr_charact_aux) + show "sat_assertion [] (insert_at 0 (fst \, snd \) []) bsyn {}" + by (metis assms(1) assms(4) get_bsem_def get_bsem_is_bsem insert_at.simps(1) prod.collapse) + show "wf_assertion_aux 0 (Suc (length [])) bsyn" + by (simp add: assms(5)) + qed (simp) + then show ?thesis + by (metis assms(1) assms(2) get_bsem_def get_bsem_is_bsem insert_at.simps(1) prod.collapse) +qed + + +lemma free_vars_charact: + assumes "wr C \ fv bsyn = {}" + and "same_syn_sem_all bsyn bsem" + and "wf_assertion_aux 0 (Suc 0) bsyn" + shows "sem C (Set.filter bsem S) = Set.filter bsem (sem C S)" (is "?A = ?B") +proof + show "?A \ ?B" + proof + fix x assume asm0: "x \ sem C (Set.filter bsem S)" + then obtain \ where "(fst x, \) \ Set.filter bsem S" "single_sem C \ (snd x)" + by (meson in_sem) + then have "agree_on (fv bsyn) \ (snd x)" + using assms(1) wr_charact by blast + then show "x \ Set.filter bsem (sem C S)" + by (metis \(fst x, \) \ Set.filter bsem S\ \\C, \\ \ snd x\ assms(2) assms(3) free_vars_syn_sem fst_conv member_filter prod.collapse single_step_then_in_sem snd_conv) + qed + show "?B \ ?A" + proof + fix x assume asm0: "x \ ?B" + then obtain \ where "bsem x" "(fst x, \) \ S" "single_sem C \ (snd x)" + by (metis in_sem member_filter) + then have "agree_on (fv bsyn) \ (snd x)" + using assms(1) wr_charact by blast + then have "bsem (fst x, \)" + using \bsem x\ agree_on_sym assms(2) assms(3) free_vars_syn_sem by fastforce + then show "x \ ?A" + by (metis \(fst x, \) \ S\ \\C, \\ \ snd x\ in_sem member_filter) + qed +qed + + +lemma filter_rule_semantic: + assumes "\ {interp_assert P} C {interp_assert Q}" + and "same_syn_sem_all bsyn bsem" + and "wr C \ fv bsyn = {}" + and "wf_assertion_aux 0 (Suc 0) bsyn" + shows "\ { interp_assert (transform_assume bsyn P) } C { interp_assert (transform_assume bsyn Q) }" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "interp_assert (transform_assume bsyn P) S" + then have "sat_assertion [] [] P (Set.filter bsem S)" + using TotalLogic.transform_assume_valid assms(2) by blast + then have "sat_assertion [] [] Q (sem C (Set.filter bsem S))" + using assms(1) hyper_hoare_tripleE by blast + moreover have "sem C (Set.filter bsem S) = Set.filter bsem (sem C S)" + using assms(2) assms(3) assms(4) free_vars_charact by presburger + ultimately show "interp_assert (transform_assume bsyn Q) (sem C S)" + using TotalLogic.transform_assume_valid assms(2) by fastforce +qed + +lemma filter_rule_syntactic: + assumes "\ {interp_assert P} C {interp_assert Q}" + and "indep_of_set b" + and "wf_assertion_aux 0 1 b" + and "wr C \ fv b = {}" + shows "\ { interp_assert (transform_assume b P) } C { interp_assert (transform_assume b Q) }" + using assms(1) filter_rule_semantic + by (metis One_nat_def assms(2) assms(3) assms(4) syntactic_charact_for_bsem) + + +definition terminates where + "terminates C \ (\\. \\'. single_sem C \ \')" + +lemma terminatesI: + assumes "\\. \\'. single_sem C \ \'" + shows "terminates C" + using terminates_def assms by auto + +lemma terminates_implies_total: + assumes "\ {P} C {Q}" + and "terminates C" + shows "\TERM {P} C {Q}" + using assms(1) +proof (rule total_hyper_tripleI) + fix \ S assume asm0: "P S \ \ \ S" + then show "\\'. \C, snd \\ \ \'" + by (metis assms(2) terminates_def) +qed + +lemma terminates_seq: + assumes "terminates C1" + and "terminates C2" + shows "terminates (C1;; C2)" + by (meson SemSeq assms(1) assms(2) terminates_def) + +lemma terminates_assign: + "terminates (Assign x e)" + by (meson SemAssign terminates_def) + +lemma terminates_havoc: + "terminates (Havoc c)" + by (meson SemHavoc terminates_def) + +lemma terminates_if: + assumes "terminates C1" + and "terminates C2" + shows "terminates (If C1 C2)" + by (meson SemIf2 assms(2) terminates_def) + + +lemma rule_lframe_exist: + fixes b :: "('a \ ('lvar \ 'lval)) \ bool" + \\b takes a mapping from keys to logical states (representing the tuple), and returns a boolean\ + + assumes "\TERM {P} C {Q}" + shows "\TERM { conj P (\S. \\. (\k. \ k \ S) \ b (fst \ \)) } C { conj Q (\S. \\. (\k. \ k \ S) \ b (fst \ \)) }" + +proof (rule total_hyper_tripleI) + fix \ S + show "Logic.conj P (\S. \\. (\k. \ k \ S) \ b (fst \ \)) S \ \ \ S \ \\'. \C, snd \\ \ \'" + by (metis (mono_tags, lifting) assms conj_def total_hyper_triple_equiv) +next + show "\ {Logic.conj P (\S. \\. (\k. \ k \ S) \ b (fst \ \))} C {Logic.conj Q (\S. \\. (\k. \ k \ S) \ b (fst \ \))}" + proof (rule hyper_hoare_tripleI) + fix S :: "('lvar, 'lval, 'b, 'c) state set" + assume asm0: "Logic.conj P (\S. \\. (\k. \ k \ S) \ b (fst \ \)) S" + then obtain \ where "(\k. \ k \ S) \ b (fst \ \)" + using conj_def[of P "\S. \\. (\k. \ k \ S) \ b (fst \ \)" S] by blast + let ?\ = "\k. SOME \'. (fst (\ k), \') \ sem C S \ single_sem C (snd (\ k)) \'" + let ?\ = "\k. (fst (\ k), ?\ k)" + have r: "\k. (fst (\ k), ?\ k) \ sem C S \ single_sem C (snd (\ k)) (?\ k)" + proof - + fix k show "(fst (\ k), ?\ k) \ sem C S \ single_sem C (snd (\ k)) (?\ k)" + proof (rule someI_ex[of "\\'. (fst (\ k), \') \ sem C S \ single_sem C (snd (\ k)) \'"]) + show "\\'. (fst (\ k), \') \ sem C S \ \C, snd (\ k)\ \ \'" + by (metis (mono_tags, lifting) \(\k. \ k \ S) \ b (fst \ \)\ asm0 assms conj_def total_hyper_triple_equiv) + qed + qed + moreover have "fst \ \ = fst \ ?\" by (rule ext) simp + ultimately have "\\. (\k. \ k \ sem C S) \ b (fst \ \)" + using \(\k. \ k \ S) \ b (fst \ \)\ by force + moreover have "Q (sem C S)" + by (metis (mono_tags, lifting) asm0 assms conj_def hyper_hoare_tripleE total_hyper_triple_def) + ultimately show "Logic.conj Q (\S. \\. (\k. \ k \ S) \ b (fst \ \)) (sem C S)" + using conj_def[of Q "\S. \\. (\k. \ k \ S) \ b (fst \ \)"] by simp + qed +qed + +lemma terminates_if_then: + assumes "terminates C1" + and "terminates C2" + shows "terminates (if_then_else b C1 C2)" +proof (rule terminatesI) + fix \ + show "\\'. \if_then_else b C1 C2, \\ \ \'" + proof (cases "b \") + case True + then show ?thesis + by (metis SemAssume SemIf1 SemSeq assms(1) if_then_else_def terminates_def) + next + case False + then show ?thesis + by (metis (no_types, opaque_lifting) SemAssume SemIf2 SemSeq assms(2) if_then_else_def lnot_def terminates_def) + qed +qed + + +definition min_prop :: "(nat \ bool) \ nat" where + "min_prop P = (SOME n. P n \ (\m. m < n \ \ P m))" + +lemma min_prop_charact: + assumes "P n" + shows "P (min_prop P) \ (\m. m < (min_prop P) \ \ P m)" + unfolding min_prop_def + using min_prop_def[of P] assms exists_least_iff[of P] someI[of "\n. P n \ (\m. m < n \ \ P m)"] + by blast + + +lemma hyper_tot_set_not_empty: + assumes "\TERM {P} C {Q}" + and "P S" + and "S \ {}" + shows "sem C S \ {}" + by (metis assms(1) assms(2) assms(3) ex_in_conv total_hyper_triple_equiv) + +lemma iterate_sem_mod_updates_same: + assumes "same_mod_updates vars S S'" + shows "same_mod_updates vars (iterate_sem n C S) (iterate_sem n C S')" + using assms +proof (induct n) + case 0 + then show ?case by simp +next + case (Suc n) + then show ?case + by (simp add: sem_update_commute) +qed + + +theorem while_synchronized_tot: + assumes "wfP lt" + and "\n. not_fv_hyper t (I n)" + and "\n. \TERM {conj (conj (I n) (holds_forall b)) (e_recorded_in_t e t)} C {conj (conj (I (Suc n)) (low_exp b)) (e_smaller_than_t e t lt)}" + shows "\TERM {conj (I 0) (low_exp b)} while_cond b C {conj (exists I) (holds_forall (lnot b))}" +proof (rule total_hyper_triple_altI) + fix S assume asm0: "conj (I 0) (low_exp b) S" + let ?S = "\n. assign_exp_to_lvar_set e t (iterate_sem n (Assume b;; C) S)" + + have S_same: "\n. same_mod_updates {t} (?S n) (iterate_sem n (Assume b;; C) S)" + by (simp add: assign_exp_to_lvar_set_same_mod_updates same_mod_updates_sym) + + + have triple: "\n. \ {conj (I n) (holds_forall b)} Assume b ;; C {conj (I (Suc n)) (low_exp b)}" + proof (rule hyper_hoare_tripleI) + fix n S assume "conj (I n) (holds_forall b) S" + let ?S = "assign_exp_to_lvar_set e t S" + have "conj (I n) (holds_forall b) ?S" + by (metis \Logic.conj (I n) (holds_forall b) S\ assms(2) conj_def holds_forall_same_assign_lvar not_fv_hyper_assign_exp) + then have "conj (conj (I n) (holds_forall b)) (e_recorded_in_t e t) ?S" + by (simp add: conj_def e_recorded_in_t_if_assigned) + then have "Logic.conj (I (Suc n)) (low_exp b) (sem (Assume b ;; C) ?S)" + by (metis (mono_tags, lifting) assms(3) conj_def hyper_hoare_tripleE sem_assume_low_exp_seq(1) total_hyper_triple_def) + moreover have "same_mod_updates {t} (sem (Assume b ;; C) ?S) (sem (Assume b ;; C) S)" + by (simp add: assign_exp_to_lvar_set_same_mod_updates same_mod_updates_sym sem_update_commute) + ultimately show "conj (I (Suc n)) (low_exp b) (sem (Assume b ;; C) S)" + by (metis assms(2) conj_def low_exp_forall_same_mod_updates not_fv_hyperE) + qed + + have "\n. iterate_sem n (Assume b ;; C) S \ {} \ conj (I n) (low_exp b) (iterate_sem n (Assume b ;; C) S)" + proof - + fix n + show "iterate_sem n (Assume b ;; C) S \ {} \ conj (I n) (low_exp b) (iterate_sem n (Assume b ;; C) S)" + proof (induct n) + case 0 + then show ?case + by (simp add: asm0) + next + case (Suc n) + then show ?case + by (metis (full_types) conj_def false_then_empty_later holds_forall_empty hyper_hoare_tripleE iterate_sem.simps(2) lessI low_exp_two_cases triple) + qed + qed + + have terminates: "\n. iterate_sem n (Assume b;; C) S = {}" + proof (rule ccontr) + assume asm0: "\ (\n. iterate_sem n (Assume b ;; C) S = {})" + + let ?R = "{(x, y). lt x y}" + let ?Q = "{ e (snd \) |\ n. \ \ ?S n}" + + obtain \0 where "\0 \ ?S 0" + by (metis all_not_in_conv asm0 false_then_empty_later holds_forall_empty holds_forall_same_assign_lvar lessI) + + show False + proof (rule wfE_min) + show "wf ?R" + using assms(1) wfP_def by blast + show "e (snd \0) \ ?Q" + using \\0 \ assign_exp_to_lvar_set e t (iterate_sem 0 (Assume b ;; C) S)\ by blast + fix z assume asm1: "z \ ?Q" "(\y. (y, z) \ {(x, y). lt x y} \ y \ ?Q)" + then obtain \ n where "z = e (snd \)" "\ \ ?S n" by blast + moreover have "conj (I n) (holds_forall b) (iterate_sem n (Assume b ;; C) S)" + proof - + have "conj (I n) (low_exp b) (iterate_sem n (Assume b ;; C) S)" + using \\n. iterate_sem n (Assume b ;; C) S \ {} \ Logic.conj (I n) (low_exp b) (iterate_sem n (Assume b ;; C) S)\ asm0 by presburger + moreover have "sem (Assume b;; C) (iterate_sem n (Assume b ;; C) S) \ {}" + using asm0 iterate_sem.simps(2) by blast + ultimately show ?thesis + by (metis asm0 conj_def false_then_empty_later lessI low_exp_two_cases) + qed + then have "conj (conj (I n) (holds_forall b)) (e_recorded_in_t e t) (?S n)" + by (metis (mono_tags, lifting) assms(2) conj_def e_recorded_in_t_if_assigned holds_forall_same_assign_lvar not_fv_hyper_assign_exp) + then have "conj (conj (I (Suc n)) (low_exp b)) (e_smaller_than_t e t lt) (sem C (?S n))" + using assms(3) hyper_hoare_tripleE total_hyper_triple_def by blast + moreover obtain \' where "\C, snd \\ \ \'" + by (meson \Logic.conj (Logic.conj (I n) (holds_forall b)) (e_recorded_in_t e t) (assign_exp_to_lvar_set e t (iterate_sem n (Assume b ;; C) S))\ assms(3) calculation(2) total_hyper_triple_def) + + then have "(fst \, \') \ sem (Assume b;; C) (?S n)" + by (metis \Logic.conj (Logic.conj (I n) (holds_forall b)) (e_recorded_in_t e t) (assign_exp_to_lvar_set e t (iterate_sem n (Assume b ;; C) S))\ calculation(2) conj_def prod.collapse sem_assume_low_exp_seq(1) single_step_then_in_sem) + then have "lt (e \') z" + by (metis (no_types, lifting) \Logic.conj (Logic.conj (I n) (holds_forall b)) (e_recorded_in_t e t) (assign_exp_to_lvar_set e t (iterate_sem n (Assume b ;; C) S))\ calculation(1) calculation(2) calculation(3) conj_def e_recorded_in_t_def e_smaller_than_t_def fst_conv sem_assume_low_exp_seq(1) snd_conv) + + moreover obtain \ where "(fst \, \) \ ?S n" "single_sem (Assume b;; C) \ \'" + by (metis \(fst \, \') \ sem (Assume b ;; C) (assign_exp_to_lvar_set e t (iterate_sem n (Assume b ;; C) S))\ fst_conv in_sem snd_conv) + then obtain l where "(l, \) \ iterate_sem n (Assume b;; C) S" + using assign_exp_to_lvar_def[of e t] assign_exp_to_lvar_set_def[of e t] image_iff prod.collapse snd_conv + by fastforce + then have "(l, \') \ iterate_sem (Suc n) (Assume b;; C) S" + by (metis \\Assume b ;; C, \\ \ \'\ iterate_sem.simps(2) single_step_then_in_sem) + then have "assign_exp_to_lvar e t (l, \') \ ?S (Suc n)" + by (simp add: assign_exp_to_lvar_set_def) + then have "e \' \ ?Q" + by (metis (mono_tags, lifting) CollectI same_outside_set_lvar_assign_exp snd_conv) + ultimately show False + using asm1(2) by blast + qed + qed + + show "conj (exists I) (holds_forall (lnot b)) (sem (while_cond b C) S)" + proof - + let ?n_emp = "min_prop (\n. iterate_sem n (Assume b;; C) S = {})" + have rr: "iterate_sem ?n_emp (Assume b;; C) S = {} \ (\m {})" + using min_prop_charact terminates by fast + show ?thesis + proof (cases "?n_emp") + case 0 + then have "S = {}" + using rr by auto + then show ?thesis + by (metis Loops.exists_def asm0 conj_def holds_forall_empty sem_assume_low_exp_seq(2) sem_seq) + next + case (Suc k) + then have "iterate_sem k (Assume b;; C) S \ {}" + using lessI rr by presburger + then have "conj (I k) (low_exp b) (iterate_sem k (Assume b ;; C) S)" + by (simp add: \\n. iterate_sem n (Assume b ;; C) S \ {} \ Logic.conj (I n) (low_exp b) (iterate_sem n (Assume b ;; C) S)\) + moreover have "holds_forall (lnot b) (iterate_sem k (Assume b;; C) S)" + proof (rule ccontr) + assume asm2: "\ holds_forall (lnot b) (iterate_sem k (Assume b ;; C) S)" + then have "holds_forall b (iterate_sem k (Assume b ;; C) S)" + by (metis calculation conj_def low_exp_two_cases) + let ?S = "assign_exp_to_lvar_set e t (iterate_sem k (Assume b ;; C) S)" + have "conj (conj (I k) (holds_forall b)) (e_recorded_in_t e t) ?S" + by (metis (no_types, lifting) \holds_forall b (iterate_sem k (Assume b ;; C) S)\ assign_exp_to_lvar_set_same_mod_updates assms(2) calculation conj_def e_recorded_in_t_if_assigned holds_forall_same_mod_updates not_fv_hyperE) + moreover have "sem (Assume b) ?S = ?S" + by (metis calculation conj_def sem_assume_low_exp(1)) + ultimately have "sem (Assume b;; C) ?S \ {}" + by (metis asm2 assms(3) holds_forall_empty holds_forall_same_assign_lvar hyper_tot_set_not_empty sem_seq) + + moreover have "same_mod_updates {t} (sem (Assume b;; C) ?S) (assign_exp_to_lvar_set e t (iterate_sem ?n_emp (Assume b;; C) S))" + by (metis Suc assign_exp_to_lvar_set_def assign_exp_to_lvar_set_same_mod_updates image_empty iterate_sem.simps(2) rr same_mod_updates_sym sem_update_commute) + ultimately show False + by (metis assign_exp_to_lvar_set_def image_empty rr same_mod_updates_empty same_mod_updates_sym) + qed + then have "\n. holds_forall (lnot b) (iterate_sem n (Assume b;; C) S) \ iterate_sem n (Assume b ;; C) S \ {} + \ (\m. m < n \ \ (holds_forall (lnot b) (iterate_sem m (Assume b;; C) S) \ iterate_sem m (Assume b ;; C) S \ {}))" + using exists_least_iff[of "\n. holds_forall (lnot b) (iterate_sem n (Assume b;; C) S) \ iterate_sem n (Assume b ;; C) S \ {}"] + using \iterate_sem k (Assume b ;; C) S \ {}\ by blast + then obtain n where def_n: "holds_forall (lnot b) (iterate_sem n (Assume b;; C) S)" "iterate_sem n (Assume b ;; C) S \ {}" + "\m. m < n \ \ (holds_forall (lnot b) (iterate_sem m (Assume b;; C) S) \ iterate_sem m (Assume b ;; C) S \ {})" by blast + moreover have "(\m. m < n \ holds_forall b (iterate_sem m (Assume b;; C) S)) \ holds_forall (lnot b) (iterate_sem n (Assume b;; C) S)" + + by (metis \\n. iterate_sem n (Assume b ;; C) S \ {} \ Logic.conj (I n) (low_exp b) (iterate_sem n (Assume b ;; C) S)\ conj_def def_n(1) def_n(2) false_then_empty_later holds_forall_empty low_exp_two_cases) + then have "sem (while_cond b C) S = iterate_sem n (Assume b;; C) S" + using triple while_synchronized_case_1 asm0 by blast + ultimately show ?thesis + by (metis Loops.exists_def \\n. iterate_sem n (Assume b ;; C) S \ {} \ Logic.conj (I n) (low_exp b) (iterate_sem n (Assume b ;; C) S)\ conj_def) + qed + qed + + + + show "terminates_in (while_cond b C) S" + proof (rule terminates_in_while_loop) + show "wfP lt" + by (simp add: assms(1)) + fix \ n + assume asm1: "\ \ iterate_sem n (Assume b ;; C) S \ b (snd \)" + let ?S = "iterate_sem n (Assume b ;; C) S" + have "conj (I n) (low_exp b) ?S" + using \\n. iterate_sem n (Assume b ;; C) S \ {} \ Logic.conj (I n) (low_exp b) (iterate_sem n (Assume b ;; C) S)\ asm1 by blast + then have "conj (I n) (holds_forall b) ?S" + by (metis asm1 conj_def holds_forall_def lnot_def low_exp_two_cases) + let ?SS = "assign_exp_to_lvar_set e t ?S" + have "conj (conj (I n) (holds_forall b)) (e_recorded_in_t e t) ?SS" + by (metis (no_types, lifting) \Logic.conj (I n) (holds_forall b) (iterate_sem n (Assume b ;; C) S)\ assign_exp_to_lvar_set_same_mod_updates assms(2) conj_def e_recorded_in_t_if_assigned holds_forall_same_mod_updates not_fv_hyperE) + then have "conj (conj (I (Suc n)) (low_exp b)) (e_smaller_than_t e t lt) (sem C ?SS)" + using assms(3) hyper_hoare_tripleE total_hyper_triple_def by blast + moreover have "assign_exp_to_lvar e t \ \ ?SS" + by (simp add: asm1 assign_exp_to_lvar_set_def) + then obtain \' where "\C, snd (assign_exp_to_lvar e t \)\ \ \'" + by (meson \Logic.conj (Logic.conj (I n) (holds_forall b)) (e_recorded_in_t e t) (assign_exp_to_lvar_set e t (iterate_sem n (Assume b ;; C) S))\ assms(3) total_hyper_triple_def) + then have "(\C, snd \\ \ \') \ (fst (assign_exp_to_lvar e t \), \') \ sem C ?SS" + by (metis \assign_exp_to_lvar e t \ \ assign_exp_to_lvar_set e t (iterate_sem n (Assume b ;; C) S)\ assign_exp_to_lvar_def prod.collapse single_step_then_in_sem snd_conv) + then have "lt (e \') (fst (fst (assign_exp_to_lvar e t \), \') t)" + by (metis calculation conj_def e_smaller_than_t_def snd_conv) + then show "\\'. (\C, snd \\ \ \') \ (\ b \' \ lt (e \') (e (snd \)))" + by (metis \(\C, snd \\ \ \') \ (fst (assign_exp_to_lvar e t \), \') \ sem C (assign_exp_to_lvar_set e t (iterate_sem n (Assume b ;; C) S))\ assign_exp_to_lvar_def fst_conv fun_upd_same) + qed +qed + +lemma total_consequence_rule: + assumes "entails P P'" + and "entails Q' Q" + and "\TERM {P'} C {Q'}" + shows "\TERM {P} C {Q}" +proof (rule total_hyper_tripleI) + show "\ {P} C {Q}" + using assms(1) assms(2) assms(3) consequence_rule total_hyper_triple_def by blast + fix \ S show "P S \ \ \ S \ \\'. \C, snd \\ \ \'" + by (meson assms(1) assms(3) entailsE total_hyper_triple_def) +qed + +theorem WhileSyncTot: + assumes "wfP lt" + and "not_fv_hyper t I" + and "\TERM {conj I (\S. \\\S. b (snd \) \ fst \ t = e (snd \))} C {conj (conj I (low_exp b)) (e_smaller_than_t e t lt)}" + shows "\TERM {conj I (low_exp b)} while_cond b C {conj I (holds_forall (lnot b))}" +proof - + define I' where "I' = (\(n::nat). I)" + have "\TERM {conj (conj I (holds_forall b)) (e_recorded_in_t e t)} C {conj (conj I (low_exp b)) (e_smaller_than_t e t lt)}" + proof (rule total_consequence_rule) + show "\TERM {conj I (\S. \\\S. b (snd \) \ fst \ t = e (snd \))} C {conj (conj I (low_exp b)) (e_smaller_than_t e t lt)}" + using assms(3) by blast + show "entails (Logic.conj (Logic.conj I (holds_forall b)) (e_recorded_in_t e t)) (Logic.conj I (\S. \\\S. b (snd \) \ fst \ t = e (snd \)))" + by (simp add: conj_def e_recorded_in_t_def entails_def holds_forall_def) + qed (simp add: entails_refl) + then have "\TERM {Logic.conj (I' 0) (low_exp b)} while_cond b C {Logic.conj (Loops.exists I') (holds_forall (lnot b))}" + using while_synchronized_tot[of lt t I' b e C] I'_def assms by blast + then show ?thesis using I'_def + by (simp add: Loops.exists_def conj_def hyper_hoare_triple_def total_hyper_triple_def) +qed + + + + +lemma total_hyper_tripleE: + assumes "\TERM {P} C {Q}" + and "P S" + and "\ \ S" + shows "\\'. (fst \, \') \ sem C S \ single_sem C (snd \) \'" + by (meson assms(1) assms(2) assms(3) total_hyper_triple_equiv) + +theorem normal_while_tot: + assumes "\n. \ {P n} Assume b {Q n}" + and "\n. \TERM {conj (Q n) (e_recorded_in_t e t)} C {conj (P (Suc n)) (e_smaller_than_t e t lt)}" + and "\ {natural_partition P} Assume (lnot b) {R}" + + and "wfP lt" + and "\n. not_fv_hyper t (P n)" + and "\n. not_fv_hyper t (Q n)" + + shows "\TERM {P 0} while_cond b C {R}" +proof (rule total_hyper_triple_altI) + fix S assume asm0: "P 0 S" + have "\n. P n (iterate_sem n (Assume b;; C) S)" + proof (rule indexed_invariant_then_power) + fix n + have "\ {Q n} C {P (Suc n)}" + proof (rule hyper_hoare_tripleI) + fix S assume asm1: "Q n S" + let ?S = "assign_exp_to_lvar_set e t S" + have "conj (Q n) (e_recorded_in_t e t) ?S" + by (metis asm1 assms(6) conj_def e_recorded_in_t_if_assigned not_fv_hyper_assign_exp) + then have "conj (P (Suc n)) (e_smaller_than_t e t lt) (sem C ?S)" + using assms(2) hyper_hoare_tripleE total_hyper_triple_def by blast + then show "P (Suc n) (sem C S)" + using assign_exp_to_lvar_set_same_mod_updates[of t S e] assms(5) conj_def not_fv_hyperE same_mod_updates_sym[of "{t}"] sem_update_commute[of "{t}"] + by metis + qed + then show " \ {P n} Assume b ;; C {P (Suc n)}" + using assms(1) seq_rule total_hyper_triple_def by blast + qed (simp add: asm0) + then have "natural_partition P (\n. iterate_sem n (Assume b;; C) S)" + by (simp add: natural_partitionI) + then show "R (sem (while_cond b C) S)" + by (metis (no_types, lifting) SUP_cong assms(3) hyper_hoare_triple_def sem_seq sem_while while_cond_def) + + show "terminates_in (while_cond b C) S" + proof (rule terminates_in_while_loop) + show "wfP lt" + by (simp add: assms(4)) + fix \ n + assume asm1: "\ \ iterate_sem n (Assume b ;; C) S \ b (snd \)" + + let ?S = "assign_exp_to_lvar_set e t (iterate_sem n (Assume b;; C) S)" + let ?\ = "assign_exp_to_lvar e t \" + have "conj (P n) (e_recorded_in_t e t) ?S" + by (metis \\n. P n (iterate_sem n (Assume b ;; C) S)\ assms(5) conj_def e_recorded_in_t_if_assigned not_fv_hyper_assign_exp) + then have "conj (Q n) (e_recorded_in_t e t) (sem (Assume b) ?S)" + + by (metis (mono_tags, lifting) assms(1) assume_sem conj_def e_recorded_in_t_def hyper_hoare_tripleE member_filter) + then have "conj (P (Suc n)) (e_smaller_than_t e t lt) (sem C (sem (Assume b) ?S))" + using assms(2) hyper_hoare_tripleE total_hyper_triple_def by blast + moreover have "?\ \ (sem (Assume b) ?S)" + by (metis asm1 assign_exp_to_lvar_set_def assume_sem comp_apply image_eqI member_filter same_outside_set_lvar_assign_exp) + then obtain \' where "(fst ?\, \') \ sem C (sem (Assume b) ?S) \ \C, snd ?\\ \ \'" + using total_hyper_tripleE assms(2)[of n] + using \Logic.conj (Q n) (e_recorded_in_t e t) (sem (Assume b) (assign_exp_to_lvar_set e t (iterate_sem n (Assume b ;; C) S)))\ by blast + then have "lt (e (snd (fst ?\, \'))) (fst (fst ?\, \') t)" + by (metis calculation conj_def e_smaller_than_t_def) + ultimately show "\\'. \C, snd \\ \ \' \ (\ b \' \ lt (e \') (e (snd \)))" + by (metis \(fst (assign_exp_to_lvar e t \), \') \ sem C (sem (Assume b) (assign_exp_to_lvar_set e t (iterate_sem n (Assume b ;; C) S))) \ \C, snd (assign_exp_to_lvar e t \)\ \ \'\ assign_exp_to_lvar_def fst_conv fun_upd_same snd_conv) + qed +qed + + +definition e_smaller_than_t_weaker where + "e_smaller_than_t_weaker e t u lt S \ (\\\S. \\'\S. fst \ u = fst \' u \ lt (e (snd \)) (fst \' t))" + + +lemma exists_terminates_loop: + assumes "wfP lt" + and "\v. \ { (\S. \\\S. e (snd \) = v \ b (snd \) \ P \ S) } if_then b C { (\S. \\\S. lt (e (snd \)) v \ P \ S) }" + and "\\. \ { P \ } while_cond b C { Q \ }" + shows "\ { (\S. \\\S. P \ S) } while_cond b C { (\S. \\\S. Q \ S)}" +proof (rule hyper_hoare_tripleI) + fix S assume asm0: "\\\S. P \ S" + then obtain \ where "\ \ S" "P \ S" by blast + show "\\\sem (while_cond b C) S. Q \ (sem (while_cond b C) S)" + proof (cases "b (snd \)") + case True + + let ?R = "{(x, y). lt x y}" + let ?Q = "{ e (snd \') |\' n. \' \ iterate_sem n (if_then b C) S \ b (snd \') \ P \' (iterate_sem n (if_then b C) S) }" + + have main_res: "\n \'. \' \ iterate_sem n (if_then b C) S \ \ b (snd \') \ P \' (iterate_sem n (if_then b C) S)" + proof (rule wfE_min) + show "wf ?R" + using assms(1) wfP_def by blast + show "e (snd \) \ ?Q" + using True \P \ S\ \\ \ S\ iterate_sem.simps(1) by fastforce + fix z + assume asm1: "z \ ?Q" "\y. (y, z) \ ?R \ y \ ?Q" + then obtain n \' where "\' \ iterate_sem n (if_then b C) S" "b (snd \')" "P \' (iterate_sem n (if_then b C) S)" "z = e (snd \')" + by blast + then have "(\S. \\\S. (b (snd \) \ lt (e (snd \)) z) \ P \ S) (sem (if_then b C) ((iterate_sem n (if_then b C) S)))" + using assms(2) hyper_hoare_tripleE by blast + then obtain \'' where "(b (snd \'') \ lt (e (snd \'')) z) \ P \'' (sem (if_then b C) ((iterate_sem n (if_then b C) S)))" + "\'' \ (sem (if_then b C) ((iterate_sem n (if_then b C) S)))" + by blast + then have "\ b (snd \'')" + by (metis (mono_tags, lifting) asm1(2) case_prodI iterate_sem.simps(2) mem_Collect_eq) + then show "\n \'. \' \ iterate_sem n (if_then b C) S \ \ b (snd \') \ P \' (iterate_sem n (if_then b C) S)" + by (metis \(b (snd \'') \ lt (e (snd \'')) z) \ P \'' (sem (if_then b C) (iterate_sem n (if_then b C) S))\ \\'' \ sem (if_then b C) (iterate_sem n (if_then b C) S)\ iterate_sem.simps(2)) + qed + then obtain n \' where "\' \ iterate_sem n (if_then b C) S" "\ b (snd \')" "P \' (iterate_sem n (if_then b C) S)" + by blast + then have "\\\sem (while_cond b C) (iterate_sem n (if_then b C) S). Q \ (sem (while_cond b C) (iterate_sem n (if_then b C) S))" + using while_exists[of P b C Q] assms(3) hyper_hoare_tripleE by blast + then show "\\\sem (while_cond b C) S. Q \ (sem (while_cond b C) S)" + by (simp add: unroll_while_sem) + next + case False + then show ?thesis + using while_exists[of P b C Q] assms(3) hyper_hoare_tripleE \P \ S\ \\ \ S\ by blast + qed +qed + +definition t_closed where + "t_closed P P_inf \ (\S. converges_sets S \ (\n. P n (S n)) \ P_inf (\n. S n))" + +lemma t_closedE: + assumes "t_closed P P_inf" + and "converges_sets S" + and "\n. P n (S n)" + shows "P_inf (\n. S n)" + using TotalLogic.t_closed_def assms(1) assms(2) assms(3) by blast + + +subsection \Total version of core rules\ + +lemma total_skip_rule: + "\TERM {P} Skip {P}" + by (meson SemSkip skip_rule total_hyper_triple_def) + +lemma total_seq_rule: + assumes "\TERM {P} C1 {R}" + and "\TERM {R} C2 {Q}" + shows "\TERM {P} Seq C1 C2 {Q}" +proof (rule total_hyper_tripleI) + show "\ {P} C1 ;; C2 {Q}" + using assms(1) assms(2) seq_rule total_hyper_triple_def by blast + fix \ S assume "P S \ \ \ S" + then obtain \' where "\C1, snd \\ \ \'" "(fst \, \') \ sem C1 S" + using assms(1) total_hyper_tripleE by blast + then obtain \'' where "\C2, \'\ \ \''" + using assms + by (metis (full_types) \P S \ \ \ S\ hyper_hoare_tripleE snd_conv total_hyper_triple_def) + then show "\\''. \C1 ;; C2, snd \\ \ \''" + using SemSeq \\C1, snd \\ \ \'\ by blast +qed + +lemma total_if_rule: + assumes "\TERM {P} C1 {Q1}" + and "\TERM {P} C2 {Q2}" + shows "\TERM {P} If C1 C2 {join Q1 Q2}" +proof (rule total_hyper_tripleI) + show "\ {P} stmt.If C1 C2 {join Q1 Q2}" + using assms(1) assms(2) if_rule total_hyper_triple_equiv by blast + fix \ S assume "P S \ \ \ S" + then show "\\'. \stmt.If C1 C2, snd \\ \ \'" + using SemIf1 assms(1) total_hyper_tripleE by blast +qed + + +lemma total_rule_exists: + assumes "\x. \TERM {P x} C {Q x}" + shows "\TERM {exists P} C {exists Q}" + using total_hyper_tripleI[of "exists P" C "exists Q"] + by (metis (mono_tags, lifting) Loops.exists_def assms hyper_hoare_triple_def total_hyper_triple_def) + + +lemma total_assign_rule: + "\TERM { (\S. P { (l, \(x := e \)) |l \. (l, \) \ S }) } (Assign x e) {P}" + using total_hyper_tripleI[of _ "Assign x e" P] + using SemAssign assign_rule by fastforce + +lemma total_havoc_rule: + "\TERM { (\S. P { (l, \(x := v)) |l \ v. (l, \) \ S }) } (Havoc x) {P}" + using total_hyper_tripleI[of _ "Havoc x" P] + using SemHavoc havoc_rule by fastforce + +lemma in_semI: + assumes "\ \ S" + and "fst \ = fst \'" + and "single_sem C (snd \) (snd \')" + shows "\' \ sem C S" + using assms + by (metis (no_types, lifting) prod.collapse single_step_then_in_sem) + + +theorem normal_while_tot_stronger: + fixes P :: "nat \ ('lvar, 'lval, 'pvar, 'pval) state set \ bool" + + assumes "\n. \ {P n} Assume b {Q n}" + and "\n. \TERM {conj (Q n) (e_recorded_in_t e t)} C {conj (P (Suc n)) (e_smaller_than_t_weaker e t u lt)}" + and "\ {natural_partition P} Assume (lnot b) {R}" + + and "wfP lt" + and "\n. not_fv_hyper t (P n)" + and "\n. not_fv_hyper t (Q n)" + + and "\n. not_fv_hyper u (P n)" + and "\n. not_fv_hyper u (Q n)" + + and "(tr :: 'lval) \ fa" + and "u \ t" + + shows "\TERM {P 0} while_cond b C {R}" +proof (rule total_hyper_triple_altI) + fix S :: "('lvar, 'lval, 'pvar, 'pval) state set" + assume asm0: "P 0 S" + have "\n. P n (iterate_sem n (Assume b;; C) S)" + proof (rule indexed_invariant_then_power) + fix n + have "\ {Q n} C {P (Suc n)}" + proof (rule hyper_hoare_tripleI) + fix S assume asm1: "Q n S" + let ?S = "assign_exp_to_lvar_set e t S" + have "conj (Q n) (e_recorded_in_t e t) ?S" + by (metis asm1 assms(6) conj_def e_recorded_in_t_if_assigned not_fv_hyper_assign_exp) + then have "conj (P (Suc n)) (e_smaller_than_t_weaker e t u lt) (sem C ?S)" + using assms(2) hyper_hoare_tripleE total_hyper_triple_def by blast + then show "P (Suc n) (sem C S)" + using assign_exp_to_lvar_set_same_mod_updates[of t] assms(5) conj_def not_fv_hyperE[of t "P (Suc n)"] same_mod_updates_sym[of "{t}"] sem_update_commute[of "{t}"] + by metis + qed + then show " \ {P n} Assume b ;; C {P (Suc n)}" + using assms(1) seq_rule total_hyper_triple_def by blast + qed (simp add: asm0) + then have "natural_partition P (\n. iterate_sem n (Assume b;; C) S)" + by (simp add: natural_partitionI) + then show "R (sem (while_cond b C) S)" + using SUP_cong assms(3) hyper_hoare_triple_def[of "natural_partition P" "Assume (lnot b)" R] sem_seq sem_while while_cond_def + by metis + + + show "terminates_in (while_cond b C) S" + proof (rule terminates_in_while_loop) + show "wfP lt" + by (simp add: assms(4)) + fix \ n + assume asm1: "\ \ iterate_sem n (Assume b ;; C) S \ b (snd \)" + + let ?S = "assign_exp_to_lvar_set e t (iterate_sem n (Assume b;; C) S)" + let ?\ = "assign_exp_to_lvar e t \" + + let ?SS = "update_lvar_set u (\\'. if \' = ?\ then tr else fa) ?S" + + have SS_def: "?SS = { ((fst \')(u := if \' = ?\ then tr else fa), snd \') |\'. \' \ ?S}" + by (simp add: update_lvar_set_def) + + have same: "same_mod_updates {u} ?S ?SS" + by (meson same_update_lvar_set) + + + let ?\' = "((fst ?\)(u := tr), snd ?\)" + + have "conj (P n) (e_recorded_in_t e t) ?S" + by (metis \\n. P n (iterate_sem n (Assume b ;; C) S)\ assms(5) conj_def e_recorded_in_t_if_assigned not_fv_hyper_assign_exp) + moreover have "e_recorded_in_t e t ?SS" + proof (rule e_recorded_in_tI) + fix \' assume "\' \ ?SS" + then show "fst \' t = e (snd \')" unfolding SS_def + using assms(10) e_recorded_in_t_def e_recorded_in_t_if_assigned by fastforce + qed + + then have "conj (P n) (e_recorded_in_t e t) ?SS" + using assms(7) calculation conj_def not_fv_hyperE[of u "P n"] same + by metis + then have "conj (Q n) (e_recorded_in_t e t) (sem (Assume b) ?SS)" + using assms(1) assume_sem[of b] conj_def e_recorded_in_t_def[of e t] hyper_hoare_tripleE[of "P n" "Assume b" "Q n" + "update_lvar_set u (\\'. if \' = assign_exp_to_lvar e t \ then tr else fa) (assign_exp_to_lvar_set e t (iterate_sem n (Assume b ;; C) S))"] member_filter[of _ "b \ snd"] + by metis + then have "conj (P (Suc n)) (e_smaller_than_t_weaker e t u lt) (sem C (sem (Assume b) ?SS))" + using assms(2) hyper_hoare_tripleE total_hyper_triple_def by blast + moreover have "?\ \ (sem (Assume b) ?S)" + by (metis asm1 assign_exp_to_lvar_set_def assume_sem comp_apply image_eqI member_filter same_outside_set_lvar_assign_exp) + moreover have "?\' \ (sem (Assume b) ?SS)" + unfolding SS_def + proof (rule in_semI) + show "((fst ?\)(u := if ?\ = assign_exp_to_lvar e t \ then tr else fa), snd ?\) \ {((fst \')(u := if \' = assign_exp_to_lvar e t \ then tr else fa), snd \') |\'. \' \ assign_exp_to_lvar_set e t (iterate_sem n (Assume b ;; C) S)}" + using asm1 assign_exp_to_lvar_set_def by fastforce + show "fst ((fst (assign_exp_to_lvar e t \))(u := if assign_exp_to_lvar e t \ = assign_exp_to_lvar e t \ then tr else fa), snd (assign_exp_to_lvar e t \)) = + fst ((fst (assign_exp_to_lvar e t \))(u := tr), snd (assign_exp_to_lvar e t \))" + by presburger + show "\Assume + b, snd ((fst (assign_exp_to_lvar e t \))(u := if assign_exp_to_lvar e t \ = assign_exp_to_lvar e t \ then tr else fa), + snd (assign_exp_to_lvar e t \))\ \ snd ((fst (assign_exp_to_lvar e t \))(u := tr), snd (assign_exp_to_lvar e t \))" + by (metis SemAssume asm1 same_outside_set_lvar_assign_exp snd_conv) + qed + then obtain \' where "(fst ?\', \') \ sem C (sem (Assume b) ?SS) \ \C, snd ?\'\ \ \'" + using total_hyper_tripleE assms(2)[of n] + using \Logic.conj (Q n) (e_recorded_in_t e t) (sem (Assume b) (update_lvar_set u (\\'. if \' = assign_exp_to_lvar e t \ then tr else fa) (assign_exp_to_lvar_set e t (iterate_sem n (Assume b ;; C) S))))\ by blast + moreover obtain \' where "\' \ sem C (sem (Assume b) ?SS)" "fst (fst ?\', \') u = fst \' u \ lt (e (snd (fst ?\', \'))) (fst \' t)" + using calculation conj_def[of "P (Suc n)" "e_smaller_than_t_weaker e t u lt"] e_smaller_than_t_weaker_def[of e t u lt] + by meson + then have "fst \' u = tr" + by simp + moreover have "fst ?\ t = fst \' t \ single_sem C (snd \) (snd \')" + proof - + obtain \0 where "\0 \ sem (Assume b) ?SS" "fst \0 = fst \'" "single_sem C (snd \0) (snd \')" + by (metis (no_types, lifting) \\' \ sem C (sem (Assume b) (update_lvar_set u (\\'. if \' = assign_exp_to_lvar e t \ then tr else fa) (assign_exp_to_lvar_set e t (iterate_sem n (Assume b ;; C) S))))\ fst_conv in_sem snd_conv) + then have "\0 \ ?SS" + by (metis (no_types, lifting) assume_sem member_filter) + then obtain \0' where + "\0' \ (assign_exp_to_lvar_set e t (iterate_sem n (Assume b ;; C) S))" + "\0 = ((fst \0')(u := if \0' = assign_exp_to_lvar e t \ then tr else fa), snd \0')" + using SS_def by auto + then have "\0 = ((fst ?\)(u := tr), snd ?\)" + by (metis (full_types) \fst \0 = fst \'\ assms(9) calculation(5) fst_conv fun_upd_same) + then show ?thesis + by (metis \\C, snd \0\ \ snd \'\ \fst \0 = fst \'\ assms(10) fst_conv fun_upd_other same_outside_set_lvar_assign_exp snd_conv) + qed + ultimately show "\\'. (\C, snd \\ \ \') \ (\ b \' \ lt (e \') (e (snd \)))" + by (metis (no_types, lifting) \fst (fst ((fst (assign_exp_to_lvar e t \))(u := tr), snd (assign_exp_to_lvar e t \)), \') u = fst \' u \ lt (e (snd (fst ((fst (assign_exp_to_lvar e t \))(u := tr), snd (assign_exp_to_lvar e t \)), \'))) (fst \' t)\ assign_exp_to_lvar_def fst_conv fun_upd_same snd_conv) + qed +qed + + + + +end \ No newline at end of file diff --git a/thys/HyperHoareLogic/document/root.bib b/thys/HyperHoareLogic/document/root.bib --- a/thys/HyperHoareLogic/document/root.bib +++ b/thys/HyperHoareLogic/document/root.bib @@ -1,684 +1,704 @@ @article{FloydLogic, author={Floyd, Robert W.}, title={Assigning Meanings to Programs}, journal={Proceedings of Symposium in Applied Mathematics}, editor={Schwartz, J. T.}, year={1967}, pages={19--32} } @article{HoareLogic, author = {Hoare, C. A. R.}, title = {An Axiomatic Basis for Computer Programming}, year = {1969}, issue_date = {Oct. 1969}, publisher = {Association for Computing Machinery}, address = {New York, NY, USA}, volume = {12}, number = {10}, issn = {0001-0782}, doi = {10.1145/363235.363259}, abstract = {In this paper an attempt is made to explore the logical foundations of computer programming by use of techniques which were first applied in the study of geometry and have later been extended to other branches of mathematics. This involves the elucidation of sets of axioms and rules of inference which can be used in proofs of the properties of computer programs. Examples are given of such axioms and rules, and a formal proof of a simple theorem is displayed. Finally, it is argued that important advantage, both theoretical and practical, may follow from a pursuance of these topics.}, journal = {Commun. ACM}, month = {oct}, pages = {576–580}, numpages = {5}, keywords = {machine-independent programming, program documentation, axiomatic method, theory of programming' proofs of programs, programming language design, formal language definition} } @INPROCEEDINGS{hyperproperties, author={Clarkson, Michael R. and Schneider, Fred B.}, booktitle={21st IEEE Computer Security Foundations Symposium}, title={Hyperproperties}, year={2008}, volume={}, number={}, pages={51-65}, doi={10.1109/CSF.2008.7}} @book{Isabelle, author = {Nipkow, Tobias and Wenzel, Markus and Paulson, Lawrence C.}, title = {Isabelle/HOL: A Proof Assistant for Higher-Order Logic}, year = {2002}, isbn = {3540433767}, publisher = {Springer-Verlag}, address = {Berlin, Heidelberg} } @inproceedings{Benton04, author = {Benton, Nick}, title = {Simple Relational Correctness Proofs for Static Analyses and Program Transformations}, year = {2004}, isbn = {158113729X}, publisher = {Association for Computing Machinery}, address = {New York, NY, USA}, doi = {10.1145/964001.964003}, abstract = {We show how some classical static analyses for imperative programs, and the optimizing transformations which they enable, may be expressed and proved correct using elementary logical and denotationaltechniques. The key ingredients are an interpretation of program properties as relations, rather than predicates, and a realization that although many program analyses are traditionally formulated in very intensional terms, the associated transformations are actually enabled by more liberal extensional properties.We illustrate our approach with formal systems for analysing and transforming while-programs. The first is a simple type system which tracks constancy and dependency information and can be used to perform dead-code elimination, constant propagation and program slicing as well as capturing a form of secure information flow. The second is a relational version of Hoare logic, which significantly generalizes our first type system and can also justify optimizations including hoisting loop invariants. Finally we show how a simple available expression analysis and redundancy elimination transformation may be justified by translation into relational Hoare logic.}, booktitle = {Proceedings of the 31st ACM SIGPLAN-SIGACT Symposium on Principles of Programming Languages}, pages = {14–25}, numpages = {12}, keywords = {dependency, partial equivalence relations, security, optimizing compilation, program analysis, denotational semantics, information flow, Hoare logic, types}, location = {Venice, Italy}, series = {POPL '04} } @inproceedings{CHL16, author = {Sousa, Marcelo and Dillig, Isil}, title = {Cartesian {H}oare {L}ogic for Verifying K-Safety Properties}, year = {2016}, isbn = {9781450342612}, publisher = {Association for Computing Machinery}, address = {New York, NY, USA}, doi = {10.1145/2908080.2908092}, abstract = {Unlike safety properties which require the absence of a “bad” program trace, k-safety properties stipulate the absence of a “bad” interaction between k traces. Examples of k-safety properties include transitivity, associativity, anti-symmetry, and monotonicity. This paper presents a sound and relatively complete calculus, called Cartesian Hoare Logic (CHL), for verifying k-safety properties. We also present an automated verification algorithm based on CHL and implement it in a tool called DESCARTES. We use DESCARTES to analyze user-defined relational operators in Java and demonstrate that DESCARTES is effective at verifying (or finding violations of) multiple k-safety properties.}, booktitle = {Proceedings of the 37th ACM SIGPLAN Conference on Programming Language Design and Implementation}, pages = {57–69}, numpages = {13}, location = {Santa Barbara, CA, USA}, series = {PLDI '16} } @article{IncorrectnessLogic, author = {O'Hearn, Peter W.}, title = {Incorrectness {L}ogic}, year = {2019}, issue_date = {January 2020}, publisher = {Association for Computing Machinery}, address = {New York, NY, USA}, volume = {4}, number = {POPL}, doi = {10.1145/3371078}, journal = {Proc. ACM Program. Lang.}, month = {dec}, articleno = {10}, numpages = {32}, keywords = {none} } @InProceedings{ReverseHL, author="de Vries, Edsko and Koutavas, Vasileios", editor="Barthe, Gilles and Pardo, Alberto and Schneider, Gerardo", title="Reverse {H}oare {L}ogic", booktitle="Software Engineering and Formal Methods", year="2011", pages="155--171", isbn="978-3-642-24690-6" } @article{InsecurityLogic, doi = {10.48550/ARXIV.2003.04791}, url = {https://arxiv.org/abs/2003.04791}, author = {Murray, Toby}, keywords = {Logic in Computer Science (cs.LO), FOS: Computer and information sciences, FOS: Computer and information sciences}, title = {An Under-Approximate Relational Logic: Heralding Logics of Insecurity, Incorrect Implementation and More}, publisher = {arXiv}, year = {2020}, journal={arXiv preprint arXiv:2003.04791}, copyright = {arXiv.org perpetual, non-exclusive license} } @article{OutcomeLogic, author = {Zilberstein, Noam and Dreyer, Derek and Silva, Alexandra}, title = {Outcome {L}ogic: A Unifying Foundation for Correctness and Incorrectness Reasoning}, year = {2023}, issue_date = {April 2023}, publisher = {Association for Computing Machinery}, address = {New York, NY, USA}, volume = {7}, number = {OOPSLA1}, doi = {10.1145/3586045}, journal = {Proc. ACM Program. Lang.}, month = {april}, articleno = {93} } @inproceedings{RHLE, author = {Dickerson, Robert and Ye, Qianchuan and Zhang, Michael K. and Delaware, Benjamin}, title = {{RHLE}: Modular Deductive Verification of Relational $\forall{}\exists{}$ Properties}, year = {2022}, isbn = {978-3-031-21036-5}, doi = {10.1007/978-3-031-21037-2\_4}, booktitle = {Programming Languages and Systems: 20th Asian Symposium, APLAS 2022, Auckland, New Zealand, December 5, 2022, Proceedings}, pages = {67–87}, numpages = {21}, location = {Auckland, New Zealand} } @article{refinement, title = {The existence of refinement mappings}, journal = {Theoretical Computer Science}, volume = {82}, number = {2}, pages = {253-284}, year = {1991}, issn = {0304-3975}, doi = {https://doi.org/10.1016/0304-3975(91)90224-P}, author = {Martín Abadi and Leslie Lamport} } @unpublished{Boogie, author = {Leino, K. Rustan M.}, title = {This is {B}oogie 2}, year = {2008}, month = {June}, url = {https://www.microsoft.com/en-us/research/publication/this-is-boogie-2-2/}, } @InProceedings{Z3, author="de Moura, Leonardo and Bj{\o}rner, Nikolaj", editor="Ramakrishnan, C. R. and Rehof, Jakob", title="Z3: An Efficient SMT Solver", booktitle="Tools and Algorithms for the Construction and Analysis of Systems", year="2008", pages="337--340", isbn="978-3-540-78800-3" } @INPROCEEDINGS{SeparationLogic, author={Reynolds, J.C.}, booktitle={Proceedings 17th Annual IEEE Symposium on Logic in Computer Science}, title={Separation logic: a logic for shared mutable data structures}, year={2002}, volume={}, number={}, pages={55-74}, doi={10.1109/LICS.2002.1029817} } @INPROCEEDINGS{SIF, author={Gray, J.W.}, booktitle={Proceedings. 1991 IEEE Computer Society Symposium on Research in Security and Privacy}, title={Toward a mathematical foundation for information flow security}, year={1991}, volume={}, number={}, pages={21-34}, keywords={}, doi={10.1109/RISP.1991.130769}, ISSN={}, month={May} } @article{CSL, title = {Resources, concurrency, and local reasoning}, journal = {Theoretical Computer Science}, volume = {375}, number = {1}, pages = {271-307}, year = {2007}, issn = {0304-3975}, doi = {https://doi.org/10.1016/j.tcs.2006.12.035}, author = {Peter W. O’Hearn}, keywords = {Concurrency, Logics of programs, Separation logic} } @article{Dijkstra75, author = {Dijkstra, Edsger W.}, title = {Guarded Commands, Nondeterminacy and Formal Derivation of Programs}, year = {1975}, issue_date = {Aug. 1975}, publisher = {Association for Computing Machinery}, address = {New York, NY, USA}, volume = {18}, number = {8}, issn = {0001-0782}, doi = {10.1145/360933.360975}, journal = {Commun. ACM}, month = {aug}, pages = {453–457}, numpages = {5}, keywords = {program semantics, programming methodology, case-construction, correctness proof, termination, nondeterminancy, programming languages, sequencing primitives, repetition, derivation of programs, programming language semantics} } @article{HypersafetyCompositionally, author = {D’Osualdo, Emanuele and Farzan, Azadeh and Dreyer, Derek}, title = {Proving Hypersafety Compositionally}, year = {2022}, issue_date = {October 2022}, publisher = {Association for Computing Machinery}, address = {New York, NY, USA}, volume = {6}, number = {OOPSLA2}, doi = {10.1145/3563298}, journal = {Proc. ACM Program. Lang.}, month = {oct}, articleno = {135}, numpages = {26}, keywords = {Modularity, Compositionality, Weakest Precondition, Hyperproperties} } @article{RelationalSL, title = {Relational separation logic}, journal = {Theoretical Computer Science}, volume = {375}, number = {1}, pages = {308-334}, year = {2007}, issn = {0304-3975}, doi = {https://doi.org/10.1016/j.tcs.2006.12.036}, author = {Hongseok Yang}, keywords = {Separation logic, Program verification, Relational reasoning, Schorr–Waite graph marking algorithm}, } @InProceedings{SecCSL, author="Ernst, Gidon and Murray, Toby", editor="Dillig, Isil and Tasiran, Serdar", title="{SecCSL}: Security Concurrent Separation Logic", booktitle="Computer Aided Verification", year="2019", address="Cham", pages="208--230", isbn="978-3-030-25543-5" } @misc{CommCSL, doi = {10.48550/ARXIV.2211.08459}, url = {https://arxiv.org/abs/2211.08459}, author = {Eilers, Marco and Dardinier, Thibault and Müller, Peter}, keywords = {Cryptography and Security (cs.CR), Programming Languages (cs.PL), FOS: Computer and information sciences, FOS: Computer and information sciences}, title = {{CommCSL}: Proving Information Flow Security for Concurrent Programs using Abstract Commutativity}, publisher = {arXiv}, year = {2022}, copyright = {arXiv.org perpetual, non-exclusive license} } @InProceedings{ISL, author="Raad, Azalea and Berdine, Josh and Dang, Hoang-Hai and Dreyer, Derek and O'Hearn, Peter and Villard, Jules", title="Local Reasoning About the Presence of Bugs: {I}ncorrectness {S}eparation {L}ogic", booktitle="Computer Aided Verification", year="2020", address="Cham", pages="225--252", isbn="978-3-030-53291-8" } @article{CISL, author = {Raad, Azalea and Berdine, Josh and Dreyer, Derek and O'Hearn, Peter W.}, title = {Concurrent {I}ncorrectness {S}eparation {L}ogic}, year = {2022}, issue_date = {January 2022}, publisher = {Association for Computing Machinery}, address = {New York, NY, USA}, volume = {6}, number = {POPL}, doi = {10.1145/3498695}, abstract = {Incorrectness separation logic (ISL) was recently introduced as a theory of under-approximate reasoning, with the goal of proving that compositional bug catchers find actual bugs. However, ISL only considers sequential programs. Here, we develop concurrent incorrectness separation logic (CISL), which extends ISL to account for bug catching in concurrent programs. Inspired by the work on Views, we design CISL as a parametric framework, which can be instantiated for a number of bug catching scenarios, including race detection, deadlock detection, and memory safety error detection. For each instance, the CISL meta-theory ensures the soundness of incorrectness reasoning for free, thereby guaranteeing that the bugs detected are true positives.}, journal = {Proc. ACM Program. Lang.}, month = {jan}, articleno = {34}, numpages = {29}, keywords = {bug catching, Concurrency, program logics, separation logic} } @InProceedings{SurveyRHL, author="Naumann, David A.", editor="Margaria, Tiziana and Steffen, Bernhard", title="Thirty-Seven Years of {R}elational {H}oare {L}ogic: Remarks on Its Principles and History", booktitle="Leveraging Applications of Formal Methods, Verification and Validation: Engineering Principles", year="2020", address="Cham", pages="93--116", abstract="Relational Hoare logics extend the applicability of modular, deductive verification to encompass important 2-run properties including dependency requirements such as confidentiality and program relations such as equivalence or similarity between program versions. A considerable number of recent works introduce different relational Hoare logics without yet converging on a core set of proof rules. This paper looks backwards to little known early work. This brings to light some principles that clarify and organize the rules as well as suggesting a new rule and a new notion of completeness.", isbn="978-3-030-61470-6" } @article{auxVariables, author = {Kleymann, Thomas}, title = {Hoare Logic and Auxiliary Variables}, year = {1999}, issue_date = {Dec 1999}, volume = {11}, number = {5}, issn = {0934-5043}, doi = {10.1007/s001650050057}, abstract = {Auxiliary variables are essential for specifying programs in Hoare Logic. They are required to relate the value of variables in different states. However, the axioms and rules of Hoare Logic turn a blind eye to the role of auxiliary variables. We stipulate a new structural rule for adjusting auxiliary variables when strengthening preconditions and weakening postconditions. Courtesy of this new rule, Hoare Logic is adaptation complete, which benefits software re-use. This property is responsible for a number of improvements. Relative completeness follows uniformly from the Most General Formula property. Moreover, one can show that Hoare Logic subsumes Vienna Development Method's (VDM) operation decomposition rules in that every derivation in VDM can be naturally embedded in Hoare Logic. Furthermore, the new treatment leads to a significant simplification in the presentation for verification calculi dealing with more interesting features such as recursion.}, journal = {Form. Asp. Comput.}, month = {dec}, pages = {541–566}, numpages = {26}, keywords = {Keywords: Hoare Logic; Auxiliary variables; Adaptation Completeness; Most General Formula; VDM} } @article{francez1983product, title={Product properties and their direct verification}, author={Francez, Nissim}, journal={Acta informatica}, volume={20}, number={4}, pages={329--344}, year={1983} } @article{aguirre2017relational, title={A relational logic for higher-order programs}, author={Aguirre, Alejandro and Barthe, Gilles and Gaboardi, Marco and Garg, Deepak and Strub, Pierre-Yves}, journal={Proceedings of the ACM on Programming Languages}, volume={1}, number={ICFP}, pages={1--29}, year={2017}, publisher={ACM New York, NY, USA} } @article{next700RHL, author = {Maillard, Kenji and Hri\c{t}cu, C\u{a}t\u{a}lin and Rivas, Exequiel and Van Muylder, Antoine}, title = {The next 700 Relational Program Logics}, year = {2019}, issue_date = {January 2020}, publisher = {Association for Computing Machinery}, address = {New York, NY, USA}, volume = {4}, number = {POPL}, doi = {10.1145/3371072}, abstract = {We propose the first framework for defining relational program logics for arbitrary monadic effects. The framework is embedded within a relational dependent type theory and is highly expressive. At the semantic level, we provide an algebraic presentation of relational specifications as a class of relative monads, and link computations and specifications by introducing relational effect observations, which map pairs of monadic computations to relational specifications in a way that respects the algebraic structure. For an arbitrary relational effect observation, we generically define the core of a sound relational program logic, and explain how to complete it to a full-fledged logic for the monadic effect at hand. We show that this generic framework can be used to define relational program logics for effects as diverse as state, input-output, nondeterminism, and discrete probabilities. We, moreover, show that by instantiating our framework with state and unbounded iteration we can embed a variant of Benton's Relational Hoare Logic, and also sketch how to reconstruct Relational Hoare Type Theory. Finally, we identify and overcome conceptual challenges that prevented previous relational program logics from properly dealing with control effects, and are the first to provide a relational program logic for exceptions.}, journal = {Proc. ACM Program. Lang.}, month = {dec}, articleno = {4}, numpages = {33}, keywords = {side-effects, monads, program verification, dependent types, foundations} } @article{AmtoftSIF, author = {Amtoft, Torben and Bandhakavi, Sruthi and Banerjee, Anindya}, title = {A Logic for Information Flow in Object-Oriented Programs}, year = {2006}, issue_date = {January 2006}, publisher = {Association for Computing Machinery}, address = {New York, NY, USA}, volume = {41}, number = {1}, issn = {0362-1340}, doi = {10.1145/1111320.1111046}, abstract = {This paper specifies, via a Hoare-like logic, an interprocedural and flow sensitive (but termination insensitive) information flow analysis for object-oriented programs. Pointer aliasing is ubiquitous in such programs, and can potentially leak confidential information. Thus the logic employs independence assertions to describe the noninterference property that formalizes confidentiality, and employs region assertions to describe possible aliasing. Programmer assertions, in the style of JML, are also allowed, thereby permitting a more fine-grained specification of information flow policy.The logic supports local reasoning about state in the style of separation logic. Small specifications are used; they mention only the variables and addresses relevant to a command. Specifications are combined using a frame rule. An algorithm for the computation of postconditions is described: under certain assumptions, there exists a strongest postcondition which the algorithm computes.}, journal = {SIGPLAN Not.}, month = {jan}, pages = {91–102}, numpages = {12}, keywords = {confidentiality, aliasing, information flow} } @InProceedings{Costanzo2014, author="Costanzo, David and Shao, Zhong", editor="Abadi, Mart{\'i}n and Kremer, Steve", title="A Separation Logic for Enforcing Declarative Information Flow Control Policies", booktitle="Principles of Security and Trust", year="2014", pages="179--198", abstract="In this paper, we present a program logic for proving that a program does not release information about sensitive data in an unintended way. The most important feature of the logic is that it provides a formal security guarantee while supporting ``declassification policies'' that describe precise conditions under which a piece of sensitive data can be released. We leverage the power of Hoare Logic to express the policies and security guarantee in terms of state predicates. This allows our system to be far more specific regarding declassification conditions than most other information flow systems.", isbn="978-3-642-54792-8" } @article{volpano1996nonInterference, title={A sound type system for secure flow analysis}, author={Volpano, Dennis and Irvine, Cynthia and Smith, Geoffrey}, journal={Journal of computer security}, volume={4}, number={2-3}, pages={167--187}, year={1996}, publisher={IOS Press} } @InProceedings{UnifyingHLandIL, author="M{\"o}ller, Bernhard and O'Hearn, Peter and Hoare, Tony", editor="Fahrenberg, Uli and Gehrke, Mai and Santocanale, Luigi and Winter, Michael", title="On Algebra of Program Correctness and Incorrectness", booktitle="Relational and Algebraic Methods in Computer Science", year="2021", address="Cham", pages="325--343", abstract="Variants of Kleene algebra have been used to provide foundations of reasoning about programs, for instance by representing Hoare Logic (HL) in algebra. That work has generally emphasised program correctness, i.e., proving the absence of bugs. Recently, Incorrectness Logic (IL) has been advanced as a formalism for the dual problem: proving the presence of bugs. IL is intended to underpin the use of logic in program testing and static bug finding. Here, we use a Kleene algebra with diamond operators and countable joins of tests, which embeds IL, and which also is complete for reasoning about the image of the embedding. Next to embedding IL, the algebra is able to embed HL, and allows making connections between IL and HL specifications. In this sense, it unifies correctness and incorrectness reasoning in one formalism.", isbn="978-3-030-88701-8" } @article{maksimovic2022exact, title={Exact Separation Logic}, author={Maksimovi{\'c}, Petar and Cronj{\"a}ger, Caroline and Sutherland, Julian and L{\"o}{\"o}w, Andreas and Ayoun, Sacha-{\'E}lie and Gardner, Philippa}, journal={arXiv preprint arXiv:2208.07200}, year={2022} } @INBOOK{DynamicLogic, author={Harel, David and Kozen, Dexter and Tiuryn, Jerzy}, booktitle={Dynamic Logic}, title={First-Order Dynamic Logic}, year={2000}, volume={}, number={}, pages={283-299}, doi={}} @book{harel1979first, title={First-order dynamic logic}, author={Harel, David}, year={1979}, publisher={Springer} } @article{hyperpropertiesDL, title={Propositional dynamic logic for hyperproperties}, author={Gutsfeld, Jens Oliver and M{\"u}ller-Olm, Markus and Ohrem, Christoph}, journal={arXiv preprint arXiv:1910.10546}, year={2019} } @article{probaHL, title={Verifying probabilistic programs using a {H}oare like logic}, author={Den Hartog, JI and de Vink, Erik P}, journal={International journal of foundations of computer science}, volume={13}, number={03}, pages={315--340}, year={2002}, publisher={World Scientific} } @inproceedings{corin2006probabilistic, title={A probabilistic {H}oare-style logic for game-based cryptographic proofs}, author={Corin, Ricardo and Den Hartog, Jerry}, booktitle={International Colloquium on Automata, Languages, and Programming}, pages={252--263}, year={2006} } @inproceedings{barthe2009formal, title={Formal certification of code-based cryptographic proofs}, author={Barthe, Gilles and Gr{\'e}goire, Benjamin and Zanella B{\'e}guelin, Santiago}, booktitle={Proceedings of the 36th annual ACM SIGPLAN-SIGACT symposium on Principles of programming languages}, pages={90--101}, year={2009} } @inproceedings{terauchi2005secure, title={Secure information flow as a safety problem}, author={Terauchi, Tachio and Aiken, Alex}, booktitle={International Static Analysis Symposium}, pages={352--367}, year={2005} } @article{selfComposition, title={Secure information flow by self-composition}, author={Barthe, Gilles and D'argenio, Pedro R and Rezk, Tamara}, journal={Mathematical Structures in Computer Science}, volume={21}, number={6}, pages={1207--1252}, year={2011}, publisher={Cambridge University Press} } @inproceedings{productPrograms, title={Relational verification using product programs}, author={Barthe, Gilles and Crespo, Juan Manuel and Kunz, C{\'e}sar}, booktitle={International Symposium on Formal Methods}, pages={200--214}, year={2011} } @article{eilers2019modular, title={Modular product programs}, author={Eilers, Marco and M{\"u}ller, Peter and Hitz, Samuel}, journal={ACM Transactions on Programming Languages and Systems (TOPLAS)}, volume={42}, number={1}, pages={1--37}, year={2019}, publisher={ACM New York, NY, USA} } @inproceedings{barthe2013beyond, title={Beyond 2-safety: Asymmetric product programs for relational program verification}, author={Barthe, Gilles and Crespo, Juan Manuel and Kunz, C{\'e}sar}, booktitle={International Symposium on Logical Foundations of Computer Science}, pages={29--43}, year={2013} } @inproceedings{clarkson2014temporal, title={Temporal logics for hyperproperties}, author={Clarkson, Michael R and Finkbeiner, Bernd and Koleini, Masoud and Micinski, Kristopher K and Rabe, Markus N and S{\'a}nchez, C{\'e}sar}, booktitle={International Conference on Principles of Security and Trust}, pages={265--284}, year={2014} } @inproceedings{coenen2019verifying, title={Verifying hyperliveness}, author={Coenen, Norine and Finkbeiner, Bernd and S{\'a}nchez, C{\'e}sar and Tentrup, Leander}, booktitle={International Conference on Computer Aided Verification}, pages={121--139}, year={2019} } @InProceedings{beyondkSafety, author="Beutner, Raven and Finkbeiner, Bernd", editor="Shoham, Sharon and Vizel, Yakir", title="Software Verification of Hyperproperties Beyond k-Safety", booktitle="Computer Aided Verification", year="2022", address="Cham", pages="341--362", abstract="Temporal hyperproperties are system properties that relate multiple execution traces. For (finite-state) hardware, temporal hyperproperties are supported by model checking algorithms, and tools for general temporal logics like HyperLTL exist. For (infinite-state) software, the analysis of temporal hyperproperties has, so far, been limited to k-safety properties, i.e., properties that stipulate the absence of a bad interaction between any k traces. In this paper, we present an automated method for the verification of {\$}{\$}{\backslash}forall ^k{\backslash}exists ^l{\$}{\$}∀k∃l-safety properties in infinite-state systems. A {\$}{\$}{\backslash}forall ^k{\backslash}exists ^l{\$}{\$}∀k∃l-safety property stipulates that for any k traces, there exist l traces such that the resulting {\$}{\$}k+l{\$}{\$}k+ltraces do not interact badly. This combination of universal and existential quantification enables us to express many properties beyond k-safety, including, for example, generalized non-interference or program refinement. Our method is based on a strategy-based instantiation of existential trace quantification combined with a program reduction, both in the context of a fixed predicate abstraction. Notably, our framework allows for mutual dependence of strategy and reduction.", isbn="978-3-031-13185-1" } @article{assaf2017hypercollecting, title={Hypercollecting semantics and its application to static analysis of information flow}, author={Assaf, Mounir and Naumann, David A and Signoles, Julien and Totel, Eric and Tronel, Fr{\'e}d{\'e}ric}, journal={ACM SIGPLAN Notices}, volume={52}, number={1}, pages={874--887}, year={2017}, publisher={ACM New York, NY, USA} } @inproceedings{cousot1977abstract, title={Abstract interpretation: a unified lattice model for static analysis of programs by construction or approximation of fixpoints}, author={Cousot, Patrick and Cousot, Radhia}, booktitle={Proceedings of the 4th ACM SIGACT-SIGPLAN symposium on Principles of programming languages}, pages={238--252}, year={1977} } @inproceedings{relationalUsingTraceLogic, title={Verifying relational properties using trace logic}, author={Barthe, Gilles and Eilers, Renate and Georgiou, Pamina and Gleiss, Bernhard and Kov{\'a}cs, Laura and Maffei, Matteo}, booktitle={2019 Formal Methods in Computer Aided Design (FMCAD)}, pages={170--178}, year={2019}, organization={IEEE} } @inproceedings{mccullough1987GNI, title={Specifications for multi-level security and a hook-up}, author={McCullough, Daryl}, booktitle={1987 IEEE Symposium on Security and Privacy}, pages={161--161}, year={1987}, organization={IEEE} } @article{mclean1996formalGNI, title={A general theory of composition for a class of" possibilistic" properties}, author={McLean, John}, journal={IEEE Transactions on Software Engineering}, volume={22}, number={1}, pages={53--67}, year={1996}, publisher={IEEE} } @inproceedings{diffPrivacyHL, title={Proving differential privacy in Hoare logic}, author={Barthe, Gilles and Gaboardi, Marco and Arias, Emilio Jes{\'u}s Gallego and Hsu, Justin and Kunz, C{\'e}sar and Strub, Pierre-Yves}, booktitle={2014 IEEE 27th Computer Security Foundations Symposium}, pages={411--424}, year={2014}, organization={IEEE} } @inproceedings{Ellora, title={An assertion-based program logic for probabilistic programs}, author={Barthe, Gilles and Espitau, Thomas and Gaboardi, Marco and Gr{\'e}goire, Benjamin and Hsu, Justin and Strub, Pierre-Yves}, booktitle={European Symposium on Programming}, pages={117--144}, year={2018} } @article{probabilisticSL, author = {Barthe, Gilles and Hsu, Justin and Liao, Kevin}, title = {A Probabilistic Separation Logic}, year = {2019}, issue_date = {January 2020}, publisher = {Association for Computing Machinery}, address = {New York, NY, USA}, volume = {4}, number = {POPL}, doi = {10.1145/3371123}, abstract = {Probabilistic independence is a useful concept for describing the result of random sampling—a basic operation in all probabilistic languages—and for reasoning about groups of random variables. Nevertheless, existing verification methods handle independence poorly, if at all. We propose a probabili stic separation logic PSL, where separation models probabilistic independence. We first give a new, probabilistic model of the logic of bunched implications (BI). We then build a program logic based on these assertions, and prove soundness of the proof system. We demonstrate our logic by verifying information-theoretic security of cryptographic constructions for several well-known tasks, including private information retrieval, oblivious transfer, secure multi-party addition, and simple oblivious RAM. Our proofs reason purely in terms of high-level properties, like independence and uniformity.}, journal = {Proc. ACM Program. Lang.}, month = {dec}, articleno = {55}, numpages = {30}, keywords = {probabilistic independence, verified cryptography, separation logic} } @article{HyperHoareLogic, title={Hyper {H}oare {L}ogic: (Dis-)Proving Program Hyperproperties (extended version)}, author={Dardinier, Thibault and M{\"u}ller, Peter}, journal={arXiv preprint arXiv:2301.10037}, year={2023}, doi = {10.48550/arXiv.2301.10037}, url = {https://doi.org/10.48550/arXiv.2301.10037} } + +@article{HyperHoareLogicPLDI, +author = {Dardinier, Thibault and M\"{u}ller, Peter}, +title = {{Hyper Hoare Logic: (Dis-)Proving Program Hyperproperties}}, +year = {2024}, +issue_date = {June 2024}, +publisher = {Association for Computing Machinery}, +address = {New York, NY, USA}, +volume = {8}, +number = {PLDI}, +url = {https://doi.org/10.1145/3656437}, +doi = {10.1145/3656437}, +journal = {Proc. ACM Program. Lang.}, +month = {jun}, +articleno = {207}, +numpages = {25}, +keywords = {Hyperproperties, Program Logic, Incorrectness Logic, Non-Interference} +} + + diff --git a/thys/HyperHoareLogic/document/root.tex b/thys/HyperHoareLogic/document/root.tex --- a/thys/HyperHoareLogic/document/root.tex +++ b/thys/HyperHoareLogic/document/root.tex @@ -1,78 +1,80 @@ \documentclass[11pt,a4paper]{article} \usepackage[T1]{fontenc} \usepackage{isabelle,isabellesym} % further packages required for unusual symbols (see also % isabellesym.sty), use only when needed %\usepackage{amssymb} %for \, \, \, \, \, \, %\, \, \, \, \, %\, \, \ %\usepackage{eurosym} %for \ %\usepackage[only,bigsqcap,bigparallel,fatsemi,interleave,sslash]{stmaryrd} %for \, \, \, \, \ %\usepackage{eufrak} %for \ ... \, \ ... \ (also included in amssymb) %\usepackage{textcomp} %for \, \, \, \, \, %\ % this should be the last package used \usepackage{pdfsetup} % urls in roman style, theory text in math-similar italics \urlstyle{rm} \isabellestyle{it} % for uniform font size %\renewcommand{\isastyle}{\isastyleminor} \begin{document} \title{Formalization of Hyper Hoare Logic:\\A Logic to (Dis-)Prove Program Hyperproperties} \author{Thibault Dardinier\\ Department of Computer Science\\ETH Zurich, Switzerland} \maketitle \begin{abstract} Hoare logics~\cite{FloydLogic, HoareLogic} are proof systems that allow one to formally establish properties of computer programs. - Traditional Hoare logics prove properties of individual program executions (so-called trace properties, such as functional correctness). + Traditional Hoare logics prove properties of individual program executions (such as functional correctness). On the one hand, Hoare logic has been generalized to prove properties of multiple executions of a program (so-called hyperproperties~\cite{hyperproperties}, such as determinism or non-interference). These program logics prove the absence of (bad combinations of) executions. On the other hand, program logics similar to Hoare logic have been proposed to disprove program properties (e.g., Incorrectness Logic~\cite{IncorrectnessLogic}), by proving the existence of (bad combinations of) executions. All of these logics have in common that they specify program properties using assertions over a fixed number of states, for instance, a single pre- and post-state for functional properties or pairs of pre- and post-states for non-interference. - In this entry, we formalize Hyper Hoare Logic~\cite{HyperHoareLogic}, a generalization of Hoare logic that lifts assertions to properties of arbitrary sets of states. - The resulting logic is simple yet expressive: its judgments can express arbitrary trace- and hyperproperties over the terminating executions of a program. - By allowing assertions to reason about sets of states, Hyper Hoare Logic can reason about both the absence and the existence of (combinations of) executions, and, thereby, supports both proving and disproving program (hyper-)properties within the same logic. - In fact, we prove that Hyper Hoare Logic subsumes the properties handled by numerous existing correctness and incorrectness logics, and can express hyperproperties that no existing Hoare logic can. - We also prove that Hyper Hoare Logic is sound and complete. + In this entry, we formalize Hyper Hoare Logic~\cite{HyperHoareLogic,HyperHoareLogicPLDI}, a generalization of Hoare logic that lifts assertions to properties of arbitrary sets of states. + The resulting logic is simple yet expressive: its judgments can express arbitrary program hyperproperties, a particular class of hyperproperties over the set of terminating executions of a program (including properties of individual program executions). + By allowing assertions to reason about sets of states, Hyper Hoare Logic can reason about both the absence and the existence of (combinations of) executions, and, thereby, supports both proving and disproving program (hyper-)properties within the same logic, including hyperproperties that no existing Hoare logic can express. + We prove that Hyper Hoare Logic is sound and complete, and demonstrate that it captures important proof principles naturally. \end{abstract} \clearpage \tableofcontents % sane default for proof documents \parindent 0pt\parskip 0.5ex % generated text of all theories \input{session} % optional bibliography -\bibliographystyle{plain} +\bibliographystyle{plainurl} +%\citestyle{alphabetic} +%\bibliographystyle{acm} +%\bibliographystyle{plain} \bibliography{root} \end{document} %%% Local Variables: %%% mode: latex %%% TeX-master: t %%% End: