| author | nipkow | 
| Tue, 06 Aug 2024 12:31:42 +0200 | |
| changeset 80643 | 11b8f2e4c3d2 | 
| parent 76216 | 9fc34f76b4e8 | 
| child 80917 | 2a77bc3b4eac | 
| permissions | -rw-r--r-- | 
| 11479 | 1 | (* Title: ZF/UNITY/Guar.thy | 
| 2 | Author: Sidi O Ehmety, Computer Laboratory | |
| 3 | Copyright 2001 University of Cambridge | |
| 4 | ||
| 5 | Guarantees, etc. | |
| 6 | ||
| 7 | From Chandy and Sanders, "Reasoning About Program Composition", | |
| 8 | Technical Report 2000-003, University of Florida, 2000. | |
| 9 | ||
| 10 | Revised by Sidi Ehmety on January 2001 | |
| 11 | ||
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 12 | Added \<in> Compatibility, weakest guarantees, etc. | 
| 11479 | 13 | |
| 14 | and Weakest existential property, | |
| 15 | from Charpentier and Chandy "Theorems about Composition", | |
| 16 | Fifth International Conference on Mathematics of Program, 2000. | |
| 17 | ||
| 18 | Theory ported from HOL. | |
| 19 | *) | |
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 20 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 21 | |
| 60770 | 22 | section\<open>The Chandy-Sanders Guarantees Operator\<close> | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 23 | |
| 16417 | 24 | theory Guar imports Comp begin | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 25 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 26 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 27 | (* To be moved to theory WFair???? *) | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 28 | lemma leadsTo_Basis': "\<lbrakk>F \<in> A co A \<union> B; F \<in> transient(A); st_set(B)\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto> B" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 29 | apply (frule constrainsD2) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 30 | apply (drule_tac B = "A-B" in constrains_weaken_L, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 31 | apply (drule_tac B = "A-B" in transient_strengthen, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 32 | apply (blast intro: ensuresI [THEN leadsTo_Basis]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 33 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 34 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 35 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 36 | (*Existential and Universal properties. We formalize the two-program | 
| 11479 | 37 | case, proving equivalence with Chandy and Sanders's n-ary definitions*) | 
| 38 | ||
| 24893 | 39 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 40 | ex_prop :: "i \<Rightarrow> o" where | 
| 76214 | 41 | "ex_prop(X) \<equiv> X<=program \<and> | 
| 61392 | 42 | (\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F \<in> X | G \<in> X \<longrightarrow> (F \<squnion> G) \<in> X)" | 
| 11479 | 43 | |
| 24893 | 44 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 45 | strict_ex_prop :: "i \<Rightarrow> o" where | 
| 76214 | 46 | "strict_ex_prop(X) \<equiv> X<=program \<and> | 
| 61392 | 47 | (\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> (F \<in> X | G \<in> X) \<longleftrightarrow> (F \<squnion> G \<in> X))" | 
| 11479 | 48 | |
| 24893 | 49 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 50 | uv_prop :: "i \<Rightarrow> o" where | 
| 76214 | 51 | "uv_prop(X) \<equiv> X<=program \<and> | 
| 52 | (SKIP \<in> X \<and> (\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F \<in> X \<and> G \<in> X \<longrightarrow> (F \<squnion> G) \<in> X))" | |
| 11479 | 53 | |
| 24893 | 54 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 55 | strict_uv_prop :: "i \<Rightarrow> o" where | 
| 76214 | 56 | "strict_uv_prop(X) \<equiv> X<=program \<and> | 
| 57 | (SKIP \<in> X \<and> (\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow>(F \<in> X \<and> G \<in> X) \<longleftrightarrow> (F \<squnion> G \<in> X)))" | |
| 11479 | 58 | |
| 24893 | 59 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 60 | guar :: "[i, i] \<Rightarrow> i" (infixl \<open>guarantees\<close> 55) where | 
| 11479 | 61 | (*higher than membership, lower than Co*) | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 62 |   "X guarantees Y \<equiv> {F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F \<squnion> G \<in> X \<longrightarrow> F \<squnion> G \<in> Y}"
 | 
| 11479 | 63 | |
| 24893 | 64 | definition | 
| 11479 | 65 | (* Weakest guarantees *) | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 66 | wg :: "[i,i] \<Rightarrow> i" where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 67 |   "wg(F,Y) \<equiv> \<Union>({X \<in> Pow(program). F:(X guarantees Y)})"
 | 
| 11479 | 68 | |
| 24893 | 69 | definition | 
| 11479 | 70 | (* Weakest existential property stronger than X *) | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 71 | wx :: "i \<Rightarrow>i" where | 
| 76214 | 72 |    "wx(X) \<equiv> \<Union>({Y \<in> Pow(program). Y<=X \<and> ex_prop(Y)})"
 | 
| 11479 | 73 | |
| 24893 | 74 | definition | 
| 61392 | 75 | (*Ill-defined programs can arise through "\<squnion>"*) | 
| 24893 | 76 | welldef :: i where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 77 |   "welldef \<equiv> {F \<in> program. Init(F) \<noteq> 0}"
 | 
| 11479 | 78 | |
| 24893 | 79 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 80 | refines :: "[i, i, i] \<Rightarrow> o" (\<open>(3_ refines _ wrt _)\<close> [10,10,10] 10) where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 81 | "G refines F wrt X \<equiv> | 
| 76214 | 82 | \<forall>H \<in> program. (F ok H \<and> G ok H \<and> F \<squnion> H \<in> welldef \<inter> X) | 
| 61392 | 83 | \<longrightarrow> (G \<squnion> H \<in> welldef \<inter> X)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 84 | |
| 24893 | 85 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 86 | iso_refines :: "[i,i, i] \<Rightarrow> o" (\<open>(3_ iso'_refines _ wrt _)\<close> [10,10,10] 10) where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 87 | "G iso_refines F wrt X \<equiv> F \<in> welldef \<inter> X \<longrightarrow> G \<in> welldef \<inter> X" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 88 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 89 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 90 | (*** existential properties ***) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 91 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 92 | lemma ex_imp_subset_program: "ex_prop(X) \<Longrightarrow> X\<subseteq>program" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 93 | by (simp add: ex_prop_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 94 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 95 | lemma ex1 [rule_format]: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 96 | "GG \<in> Fin(program) | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 97 | \<Longrightarrow> ex_prop(X) \<longrightarrow> GG \<inter> X\<noteq>0 \<longrightarrow> OK(GG, (\<lambda>G. G)) \<longrightarrow>(\<Squnion>G \<in> GG. G) \<in> X" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 98 | unfolding ex_prop_def | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 99 | apply (erule Fin_induct) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 100 | apply (simp_all add: OK_cons_iff) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 101 | apply (safe elim!: not_emptyE, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 102 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 103 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 104 | lemma ex2 [rule_format]: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 105 | "X \<subseteq> program \<Longrightarrow> | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 106 | (\<forall>GG \<in> Fin(program). GG \<inter> X \<noteq> 0 \<longrightarrow> OK(GG,(\<lambda>G. G))\<longrightarrow>(\<Squnion>G \<in> GG. G) \<in> X) | 
| 46823 | 107 | \<longrightarrow> ex_prop(X)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 108 | apply (unfold ex_prop_def, clarify) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 109 | apply (drule_tac x = "{F,G}" in bspec)
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 110 | apply (simp_all add: OK_iff_ok) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 111 | apply (auto intro: ok_sym) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 112 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 113 | |
| 76214 | 114 | (*Chandy \<and> Sanders take this as a definition*) | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 115 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 116 | lemma ex_prop_finite: | 
| 76214 | 117 | "ex_prop(X) \<longleftrightarrow> (X\<subseteq>program \<and> | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 118 | (\<forall>GG \<in> Fin(program). GG \<inter> X \<noteq> 0 \<and> OK(GG,(\<lambda>G. G))\<longrightarrow>(\<Squnion>G \<in> GG. G) \<in> X))" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 119 | apply auto | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 120 | apply (blast intro: ex1 ex2 dest: ex_imp_subset_program)+ | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 121 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 122 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 123 | (* Equivalent definition of ex_prop given at the end of section 3*) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 124 | lemma ex_prop_equiv: | 
| 46823 | 125 | "ex_prop(X) \<longleftrightarrow> | 
| 76214 | 126 | X\<subseteq>program \<and> (\<forall>G \<in> program. (G \<in> X \<longleftrightarrow> (\<forall>H \<in> program. (G component_of H) \<longrightarrow> H \<in> X)))" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 127 | apply (unfold ex_prop_def component_of_def, safe, force, force, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 128 | apply (subst Join_commute) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 129 | apply (blast intro: ok_sym) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 130 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 131 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 132 | (*** universal properties ***) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 133 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 134 | lemma uv_imp_subset_program: "uv_prop(X)\<Longrightarrow> X\<subseteq>program" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 135 | unfolding uv_prop_def | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 136 | apply (simp (no_asm_simp)) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 137 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 138 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 139 | lemma uv1 [rule_format]: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 140 | "GG \<in> Fin(program) \<Longrightarrow> | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 141 | (uv_prop(X)\<longrightarrow> GG \<subseteq> X \<and> OK(GG, (\<lambda>G. G)) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in> X)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 142 | unfolding uv_prop_def | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 143 | apply (erule Fin_induct) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 144 | apply (auto simp add: OK_cons_iff) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 145 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 146 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 147 | lemma uv2 [rule_format]: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 148 | "X\<subseteq>program \<Longrightarrow> | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 149 | (\<forall>GG \<in> Fin(program). GG \<subseteq> X \<and> OK(GG,(\<lambda>G. G)) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in> X) | 
| 46823 | 150 | \<longrightarrow> uv_prop(X)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 151 | apply (unfold uv_prop_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 152 | apply (drule_tac x = 0 in bspec, simp+) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 153 | apply (drule_tac x = "{F,G}" in bspec, simp) 
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 154 | apply (force dest: ok_sym simp add: OK_iff_ok) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 155 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 156 | |
| 76214 | 157 | (*Chandy \<and> Sanders take this as a definition*) | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 158 | lemma uv_prop_finite: | 
| 76214 | 159 | "uv_prop(X) \<longleftrightarrow> X\<subseteq>program \<and> | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 160 | (\<forall>GG \<in> Fin(program). GG \<subseteq> X \<and> OK(GG, \<lambda>G. G) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in> X)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 161 | apply auto | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 162 | apply (blast dest: uv_imp_subset_program) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 163 | apply (blast intro: uv1) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 164 | apply (blast intro!: uv2 dest:) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 165 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 166 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 167 | (*** guarantees ***) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 168 | lemma guaranteesI: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 169 | "\<lbrakk>(\<And>G. \<lbrakk>F ok G; F \<squnion> G \<in> X; G \<in> program\<rbrakk> \<Longrightarrow> F \<squnion> G \<in> Y); | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 170 | F \<in> program\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 171 | \<Longrightarrow> F \<in> X guarantees Y" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 172 | by (simp add: guar_def component_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 173 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 174 | lemma guaranteesD: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 175 | "\<lbrakk>F \<in> X guarantees Y; F ok G; F \<squnion> G \<in> X; G \<in> program\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 176 | \<Longrightarrow> F \<squnion> G \<in> Y" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 177 | by (simp add: guar_def component_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 178 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 179 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 180 | (*This version of guaranteesD matches more easily in the conclusion | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 181 | The major premise can no longer be F\<subseteq>H since we need to reason about G*) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 182 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 183 | lemma component_guaranteesD: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 184 | "\<lbrakk>F \<in> X guarantees Y; F \<squnion> G = H; H \<in> X; F ok G; G \<in> program\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 185 | \<Longrightarrow> H \<in> Y" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 186 | by (simp add: guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 187 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 188 | lemma guarantees_weaken: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 189 | "\<lbrakk>F \<in> X guarantees X'; Y \<subseteq> X; X' \<subseteq> Y'\<rbrakk> \<Longrightarrow> F \<in> Y guarantees Y'" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 190 | by (simp add: guar_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 191 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 192 | lemma subset_imp_guarantees_program: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 193 | "X \<subseteq> Y \<Longrightarrow> X guarantees Y = program" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 194 | by (unfold guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 195 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 196 | (*Equivalent to subset_imp_guarantees_UNIV but more intuitive*) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 197 | lemma subset_imp_guarantees: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 198 | "\<lbrakk>X \<subseteq> Y; F \<in> program\<rbrakk> \<Longrightarrow> F \<in> X guarantees Y" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 199 | by (unfold guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 200 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 201 | lemma component_of_Join1: "F ok G \<Longrightarrow> F component_of (F \<squnion> G)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 202 | by (unfold component_of_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 203 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 204 | lemma component_of_Join2: "F ok G \<Longrightarrow> G component_of (F \<squnion> G)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 205 | apply (subst Join_commute) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 206 | apply (blast intro: ok_sym component_of_Join1) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 207 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 208 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 209 | (*Remark at end of section 4.1 *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 210 | lemma ex_prop_imp: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 211 | "ex_prop(Y) \<Longrightarrow> (Y = (program guarantees Y))" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 212 | apply (simp (no_asm_use) add: ex_prop_equiv guar_def component_of_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 213 | apply clarify | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 214 | apply (rule equalityI, blast, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 215 | apply (drule_tac x = x in bspec, assumption, force) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 216 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 217 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 218 | lemma guarantees_imp: "(Y = program guarantees Y) \<Longrightarrow> ex_prop(Y)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 219 | unfolding guar_def | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 220 | apply (simp (no_asm_simp) add: ex_prop_equiv) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 221 | apply safe | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 222 | apply (blast intro: elim: equalityE) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 223 | apply (simp_all (no_asm_use) add: component_of_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 224 | apply (force elim: equalityE)+ | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 225 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 226 | |
| 46823 | 227 | lemma ex_prop_equiv2: "(ex_prop(Y)) \<longleftrightarrow> (Y = program guarantees Y)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 228 | by (blast intro: ex_prop_imp guarantees_imp) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 229 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 230 | (** Distributive laws. Re-orient to perform miniscoping **) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 231 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 232 | lemma guarantees_UN_left: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 233 | "i \<in> I \<Longrightarrow>(\<Union>i \<in> I. X(i)) guarantees Y = (\<Inter>i \<in> I. X(i) guarantees Y)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 234 | unfolding guar_def | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 235 | apply (rule equalityI, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 236 | prefer 2 apply force | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 237 | apply blast+ | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 238 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 239 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 240 | lemma guarantees_Un_left: | 
| 46823 | 241 | "(X \<union> Y) guarantees Z = (X guarantees Z) \<inter> (Y guarantees Z)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 242 | unfolding guar_def | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 243 | apply (rule equalityI, safe, blast+) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 244 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 245 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 246 | lemma guarantees_INT_right: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 247 | "i \<in> I \<Longrightarrow> X guarantees (\<Inter>i \<in> I. Y(i)) = (\<Inter>i \<in> I. X guarantees Y(i))" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 248 | unfolding guar_def | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 249 | apply (rule equalityI, safe, blast+) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 250 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 251 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 252 | lemma guarantees_Int_right: | 
| 46823 | 253 | "Z guarantees (X \<inter> Y) = (Z guarantees X) \<inter> (Z guarantees Y)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 254 | by (unfold guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 255 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 256 | lemma guarantees_Int_right_I: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 257 | "\<lbrakk>F \<in> Z guarantees X; F \<in> Z guarantees Y\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 258 | \<Longrightarrow> F \<in> Z guarantees (X \<inter> Y)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 259 | by (simp (no_asm_simp) add: guarantees_Int_right) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 260 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 261 | lemma guarantees_INT_right_iff: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 262 | "i \<in> I\<Longrightarrow> (F \<in> X guarantees (\<Inter>i \<in> I. Y(i))) \<longleftrightarrow> | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 263 | (\<forall>i \<in> I. F \<in> X guarantees Y(i))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 264 | by (simp add: guarantees_INT_right INT_iff, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 265 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 266 | |
| 46823 | 267 | lemma shunting: "(X guarantees Y) = (program guarantees ((program-X) \<union> Y))" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 268 | by (unfold guar_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 269 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 270 | lemma contrapositive: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 271 | "(X guarantees Y) = (program - Y) guarantees (program -X)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 272 | by (unfold guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 273 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 274 | (** The following two can be expressed using intersection and subset, which | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 275 | is more faithful to the text but looks cryptic. | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 276 | **) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 277 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 278 | lemma combining1: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 279 | "\<lbrakk>F \<in> V guarantees X; F \<in> (X \<inter> Y) guarantees Z\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 280 | \<Longrightarrow> F \<in> (V \<inter> Y) guarantees Z" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 281 | by (unfold guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 282 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 283 | lemma combining2: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 284 | "\<lbrakk>F \<in> V guarantees (X \<union> Y); F \<in> Y guarantees Z\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 285 | \<Longrightarrow> F \<in> V guarantees (X \<union> Z)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 286 | by (unfold guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 287 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 288 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 289 | (** The following two follow Chandy-Sanders, but the use of object-quantifiers | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 290 | does not suit Isabelle... **) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 291 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 292 | (*Premise should be (\<And>i. i \<in> I \<Longrightarrow> F \<in> X guarantees Y i) *) | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 293 | lemma all_guarantees: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 294 | "\<lbrakk>\<forall>i \<in> I. F \<in> X guarantees Y(i); i \<in> I\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 295 | \<Longrightarrow> F \<in> X guarantees (\<Inter>i \<in> I. Y(i))" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 296 | by (unfold guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 297 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 298 | (*Premises should be \<lbrakk>F \<in> X guarantees Y i; i \<in> I\<rbrakk> *) | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 299 | lemma ex_guarantees: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 300 | "\<exists>i \<in> I. F \<in> X guarantees Y(i) \<Longrightarrow> F \<in> X guarantees (\<Union>i \<in> I. Y(i))" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 301 | by (unfold guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 302 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 303 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 304 | (*** Additional guarantees laws, by lcp ***) | 
| 11479 | 305 | |
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 306 | lemma guarantees_Join_Int: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 307 | "\<lbrakk>F \<in> U guarantees V; G \<in> X guarantees Y; F ok G\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 308 | \<Longrightarrow> F \<squnion> G: (U \<inter> X) guarantees (V \<inter> Y)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 309 | |
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 310 | unfolding guar_def | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 311 | apply (simp (no_asm)) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 312 | apply safe | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 313 | apply (simp add: Join_assoc) | 
| 61392 | 314 | apply (subgoal_tac "F \<squnion> G \<squnion> Ga = G \<squnion> (F \<squnion> Ga) ") | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 315 | apply (simp add: ok_commute) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 316 | apply (simp (no_asm_simp) add: Join_ac) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 317 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 318 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 319 | lemma guarantees_Join_Un: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 320 | "\<lbrakk>F \<in> U guarantees V; G \<in> X guarantees Y; F ok G\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 321 | \<Longrightarrow> F \<squnion> G: (U \<union> X) guarantees (V \<union> Y)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 322 | unfolding guar_def | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 323 | apply (simp (no_asm)) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 324 | apply safe | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 325 | apply (simp add: Join_assoc) | 
| 61392 | 326 | apply (subgoal_tac "F \<squnion> G \<squnion> Ga = G \<squnion> (F \<squnion> Ga) ") | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 327 | apply (rotate_tac 4) | 
| 61392 | 328 | apply (drule_tac x = "F \<squnion> Ga" in bspec) | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 329 | apply (simp (no_asm)) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 330 | apply (force simp add: ok_commute) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 331 | apply (simp (no_asm_simp) add: Join_ac) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 332 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 333 | |
| 61392 | 334 | lemma guarantees_JOIN_INT: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 335 | "\<lbrakk>\<forall>i \<in> I. F(i) \<in> X(i) guarantees Y(i); OK(I,F); i \<in> I\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 336 | \<Longrightarrow> (\<Squnion>i \<in> I. F(i)) \<in> (\<Inter>i \<in> I. X(i)) guarantees (\<Inter>i \<in> I. Y(i))" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 337 | apply (unfold guar_def, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 338 | prefer 2 apply blast | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 339 | apply (drule_tac x = xa in bspec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 340 | apply (simp_all add: INT_iff, safe) | 
| 61392 | 341 | apply (drule_tac x = "(\<Squnion>x \<in> (I-{xa}) . F (x)) \<squnion> G" and A=program in bspec)
 | 
| 342 | apply (auto intro: OK_imp_ok simp add: Join_assoc [symmetric] JOIN_Join_diff JOIN_absorb) | |
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 343 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 344 | |
| 61392 | 345 | lemma guarantees_JOIN_UN: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 346 | "\<lbrakk>\<forall>i \<in> I. F(i) \<in> X(i) guarantees Y(i); OK(I,F)\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 347 | \<Longrightarrow> JOIN(I,F) \<in> (\<Union>i \<in> I. X(i)) guarantees (\<Union>i \<in> I. Y(i))" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 348 | apply (unfold guar_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 349 | apply (drule_tac x = y in bspec, simp_all, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 350 | apply (rename_tac G y) | 
| 61392 | 351 | apply (drule_tac x = "JOIN (I-{y}, F) \<squnion> G" and A=program in bspec)
 | 
| 352 | apply (auto intro: OK_imp_ok simp add: Join_assoc [symmetric] JOIN_Join_diff JOIN_absorb) | |
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 353 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 354 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 355 | (*** guarantees laws for breaking down the program, by lcp ***) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 356 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 357 | lemma guarantees_Join_I1: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 358 | "\<lbrakk>F \<in> X guarantees Y; F ok G\<rbrakk> \<Longrightarrow> F \<squnion> G \<in> X guarantees Y" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 359 | apply (simp add: guar_def, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 360 | apply (simp add: Join_assoc) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 361 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 362 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 363 | lemma guarantees_Join_I2: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 364 | "\<lbrakk>G \<in> X guarantees Y; F ok G\<rbrakk> \<Longrightarrow> F \<squnion> G \<in> X guarantees Y" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 365 | apply (simp add: Join_commute [of _ G] ok_commute [of _ G]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 366 | apply (blast intro: guarantees_Join_I1) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 367 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 368 | |
| 61392 | 369 | lemma guarantees_JOIN_I: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 370 | "\<lbrakk>i \<in> I; F(i) \<in> X guarantees Y; OK(I,F)\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 371 | \<Longrightarrow> (\<Squnion>i \<in> I. F(i)) \<in> X guarantees Y" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 372 | apply (unfold guar_def, safe) | 
| 61392 | 373 | apply (drule_tac x = "JOIN (I-{i},F) \<squnion> G" in bspec)
 | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 374 | apply (simp (no_asm)) | 
| 61392 | 375 | apply (auto intro: OK_imp_ok simp add: JOIN_Join_diff Join_assoc [symmetric]) | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 376 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 377 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 378 | (*** well-definedness ***) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 379 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 380 | lemma Join_welldef_D1: "F \<squnion> G \<in> welldef \<Longrightarrow> programify(F) \<in> welldef" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 381 | by (unfold welldef_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 382 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 383 | lemma Join_welldef_D2: "F \<squnion> G \<in> welldef \<Longrightarrow> programify(G) \<in> welldef" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 384 | by (unfold welldef_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 385 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 386 | (*** refinement ***) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 387 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 388 | lemma refines_refl: "F refines F wrt X" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 389 | by (unfold refines_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 390 | |
| 76214 | 391 | (* More results on guarantees, added by Sidi Ehmety from Chandy \<and> Sander, section 6 *) | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 392 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 393 | lemma wg_type: "wg(F, X) \<subseteq> program" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 394 | by (unfold wg_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 395 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 396 | lemma guarantees_type: "X guarantees Y \<subseteq> program" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 397 | by (unfold guar_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 398 | |
| 76214 | 399 | lemma wgD2: "G \<in> wg(F, X) \<Longrightarrow> G \<in> program \<and> F \<in> program" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 400 | apply (unfold wg_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 401 | apply (blast dest: guarantees_type [THEN subsetD]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 402 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 403 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 404 | lemma guarantees_equiv: | 
| 46823 | 405 | "(F \<in> X guarantees Y) \<longleftrightarrow> | 
| 76214 | 406 | F \<in> program \<and> (\<forall>H \<in> program. H \<in> X \<longrightarrow> (F component_of H \<longrightarrow> H \<in> Y))" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 407 | by (unfold guar_def component_of_def, force) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 408 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 409 | lemma wg_weakest: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 410 | "\<And>X. \<lbrakk>F \<in> (X guarantees Y); X \<subseteq> program\<rbrakk> \<Longrightarrow> X \<subseteq> wg(F,Y)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 411 | by (unfold wg_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 412 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 413 | lemma wg_guarantees: "F \<in> program \<Longrightarrow> F \<in> wg(F,Y) guarantees Y" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 414 | by (unfold wg_def guar_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 415 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 416 | lemma wg_equiv: | 
| 46823 | 417 | "H \<in> wg(F,X) \<longleftrightarrow> | 
| 76214 | 418 | ((F component_of H \<longrightarrow> H \<in> X) \<and> F \<in> program \<and> H \<in> program)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 419 | apply (simp add: wg_def guarantees_equiv) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 420 | apply (rule iffI, safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 421 | apply (rule_tac [4] x = "{H}" in bexI)
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 422 | apply (rule_tac [3] x = "{H}" in bexI, blast+)
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 423 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 424 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 425 | lemma component_of_wg: | 
| 76214 | 426 | "F component_of H \<Longrightarrow> H \<in> wg(F,X) \<longleftrightarrow> (H \<in> X \<and> F \<in> program \<and> H \<in> program)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 427 | by (simp (no_asm_simp) add: wg_equiv) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 428 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 429 | lemma wg_finite [rule_format]: | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 430 | "\<forall>FF \<in> Fin(program). FF \<inter> X \<noteq> 0 \<longrightarrow> OK(FF, \<lambda>F. F) | 
| 46823 | 431 | \<longrightarrow> (\<forall>F \<in> FF. ((\<Squnion>F \<in> FF. F) \<in> wg(F,X)) \<longleftrightarrow> ((\<Squnion>F \<in> FF. F) \<in> X))" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 432 | apply clarify | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 433 | apply (subgoal_tac "F component_of (\<Squnion>F \<in> FF. F) ") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 434 | apply (drule_tac X = X in component_of_wg) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 435 | apply (force dest!: Fin.dom_subset [THEN subsetD, THEN PowD]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 436 | apply (simp_all add: component_of_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 437 | apply (rule_tac x = "\<Squnion>F \<in> (FF-{F}) . F" in exI)
 | 
| 61392 | 438 | apply (auto intro: JOIN_Join_diff dest: ok_sym simp add: OK_iff_ok) | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 439 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 440 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 441 | lemma wg_ex_prop: | 
| 76214 | 442 | "ex_prop(X) \<Longrightarrow> (F \<in> X) \<longleftrightarrow> (\<forall>H \<in> program. H \<in> wg(F,X) \<and> F \<in> program)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 443 | apply (simp (no_asm_use) add: ex_prop_equiv wg_equiv) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 444 | apply blast | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 445 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 446 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 447 | (** From Charpentier and Chandy "Theorems About Composition" **) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 448 | (* Proposition 2 *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 449 | lemma wx_subset: "wx(X)\<subseteq>X" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 450 | by (unfold wx_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 451 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 452 | lemma wx_ex_prop: "ex_prop(wx(X))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 453 | apply (simp (no_asm_use) add: ex_prop_def wx_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 454 | apply safe | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 455 | apply blast | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 456 | apply (rule_tac x=x in bexI, force, simp)+ | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 457 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 458 | |
| 46823 | 459 | lemma wx_weakest: "\<forall>Z. Z\<subseteq>program \<longrightarrow> Z\<subseteq> X \<longrightarrow> ex_prop(Z) \<longrightarrow> Z \<subseteq> wx(X)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 460 | by (unfold wx_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 461 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 462 | (* Proposition 6 *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 463 | lemma wx'_ex_prop: | 
| 61392 | 464 |  "ex_prop({F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F \<squnion> G \<in> X})"
 | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 465 | apply (unfold ex_prop_def, safe) | 
| 61392 | 466 | apply (drule_tac x = "G \<squnion> Ga" in bspec) | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 467 | apply (simp (no_asm)) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 468 | apply (force simp add: Join_assoc) | 
| 61392 | 469 | apply (drule_tac x = "F \<squnion> Ga" in bspec) | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 470 | apply (simp (no_asm)) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 471 | apply (simp (no_asm_use)) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 472 | apply safe | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 473 | apply (simp (no_asm_simp) add: ok_commute) | 
| 61392 | 474 | apply (subgoal_tac "F \<squnion> G = G \<squnion> F") | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 475 | apply (simp (no_asm_simp) add: Join_assoc) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 476 | apply (simp (no_asm) add: Join_commute) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 477 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 478 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 479 | (* Equivalence with the other definition of wx *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 480 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 481 | lemma wx_equiv: | 
| 61392 | 482 |      "wx(X) = {F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> (F \<squnion> G) \<in> X}"
 | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 483 | unfolding wx_def | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 484 | apply (rule equalityI, safe, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 485 | apply (simp (no_asm_use) add: ex_prop_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 486 | apply blast | 
| 61392 | 487 | apply (rule_tac B = "{F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F \<squnion> G \<in> X}" 
 | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 488 | in UnionI, | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 489 | safe) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 490 | apply (rule_tac [2] wx'_ex_prop) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 491 | apply (drule_tac x=SKIP in bspec, simp)+ | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 492 | apply auto | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 493 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 494 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 495 | (* Propositions 7 to 11 are all about this second definition of wx. And | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 496 | by equivalence between the two definition, they are the same as the ones proved *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 497 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 498 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 499 | (* Proposition 12 *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 500 | (* Main result of the paper *) | 
| 46823 | 501 | lemma guarantees_wx_eq: "(X guarantees Y) = wx((program-X) \<union> Y)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 502 | by (auto simp add: guar_def wx_equiv) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 503 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 504 | (* | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 505 | {* Corollary, but this result has already been proved elsewhere *}
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 506 | "ex_prop(X guarantees Y)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 507 | *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 508 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 509 | (* Rules given in section 7 of Chandy and Sander's | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 510 | Reasoning About Program composition paper *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 511 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 512 | lemma stable_guarantees_Always: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 513 | "\<lbrakk>Init(F) \<subseteq> A; F \<in> program\<rbrakk> \<Longrightarrow> F \<in> stable(A) guarantees Always(A)" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 514 | apply (rule guaranteesI) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 515 | prefer 2 apply assumption | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 516 | apply (simp (no_asm) add: Join_commute) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 517 | apply (rule stable_Join_Always1) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 518 | apply (simp_all add: invariant_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 519 | apply (auto simp add: programify_def initially_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 520 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 521 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 522 | lemma constrains_guarantees_leadsTo: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 523 | "\<lbrakk>F \<in> transient(A); st_set(B)\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69587diff
changeset | 524 | \<Longrightarrow> F: (A co A \<union> B) guarantees (A \<longmapsto> (B-A))" | 
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 525 | apply (rule guaranteesI) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 526 | prefer 2 apply (blast dest: transient_type [THEN subsetD]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 527 | apply (rule leadsTo_Basis') | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 528 | apply (blast intro: constrains_weaken_R) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 529 | apply (blast intro!: Join_transient_I1, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
12195diff
changeset | 530 | done | 
| 11479 | 531 | |
| 532 | end |