| author | wenzelm | 
| Mon, 11 Sep 2000 17:40:41 +0200 | |
| changeset 9919 | 3cf12ab0b8ac | 
| parent 9025 | e50c0764e522 | 
| child 10064 | 1a77667b21ef | 
| permissions | -rw-r--r-- | 
| 4776 | 1 | (* Title: HOL/UNITY/UNITY | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1998 University of Cambridge | |
| 5 | ||
| 6 | The basic UNITY theory (revised version, based upon the "co" operator) | |
| 7 | ||
| 8 | From Misra, "A Logic for Concurrent Programming", 1994 | |
| 9 | *) | |
| 10 | ||
| 9000 | 11 | set timing; | 
| 4776 | 12 | |
| 8311 
6218522253e7
new mostly working version; Alloc nearly converted to "Rename"
 paulson parents: 
8216diff
changeset | 13 | (*Perhaps equalities.ML shouldn't add this in the first place!*) | 
| 
6218522253e7
new mostly working version; Alloc nearly converted to "Rename"
 paulson parents: 
8216diff
changeset | 14 | Delsimps [image_Collect]; | 
| 
6218522253e7
new mostly working version; Alloc nearly converted to "Rename"
 paulson parents: 
8216diff
changeset | 15 | |
| 4776 | 16 | |
| 6535 | 17 | (*** The abstract type of programs ***) | 
| 18 | ||
| 19 | val rep_ss = simpset() addsimps | |
| 20 | [Init_def, Acts_def, mk_program_def, Program_def, Rep_Program, | |
| 21 | Rep_Program_inverse, Abs_Program_inverse]; | |
| 22 | ||
| 23 | ||
| 24 | Goal "Id : Acts F"; | |
| 25 | by (cut_inst_tac [("x", "F")] Rep_Program 1);
 | |
| 26 | by (auto_tac (claset(), rep_ss)); | |
| 27 | qed "Id_in_Acts"; | |
| 28 | AddIffs [Id_in_Acts]; | |
| 29 | ||
| 30 | Goal "insert Id (Acts F) = Acts F"; | |
| 31 | by (simp_tac (simpset() addsimps [insert_absorb, Id_in_Acts]) 1); | |
| 32 | qed "insert_Id_Acts"; | |
| 33 | AddIffs [insert_Id_Acts]; | |
| 34 | ||
| 35 | (** Inspectors for type "program" **) | |
| 36 | ||
| 37 | Goal "Init (mk_program (init,acts)) = init"; | |
| 38 | by (auto_tac (claset(), rep_ss)); | |
| 39 | qed "Init_eq"; | |
| 40 | ||
| 41 | Goal "Acts (mk_program (init,acts)) = insert Id acts"; | |
| 42 | by (auto_tac (claset(), rep_ss)); | |
| 43 | qed "Acts_eq"; | |
| 44 | ||
| 45 | Addsimps [Acts_eq, Init_eq]; | |
| 46 | ||
| 47 | ||
| 8327 | 48 | (** Equality for UNITY programs **) | 
| 7915 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7826diff
changeset | 49 | |
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7826diff
changeset | 50 | Goal "mk_program (Init F, Acts F) = F"; | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7826diff
changeset | 51 | by (cut_inst_tac [("x", "F")] Rep_Program 1);
 | 
| 6535 | 52 | by (auto_tac (claset(), rep_ss)); | 
| 53 | by (REPEAT (dres_inst_tac [("f", "Abs_Program")] arg_cong 1));
 | |
| 7915 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7826diff
changeset | 54 | by (asm_full_simp_tac (rep_ss addsimps [insert_absorb]) 1); | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7826diff
changeset | 55 | qed "surjective_mk_program"; | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7826diff
changeset | 56 | |
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7826diff
changeset | 57 | Goal "[| Init F = Init G; Acts F = Acts G |] ==> F = G"; | 
| 8988 | 58 | by (res_inst_tac [("t", "F")] (surjective_mk_program RS subst) 1);
 | 
| 59 | by (res_inst_tac [("t", "G")] (surjective_mk_program RS subst) 1);
 | |
| 60 | by (Asm_simp_tac 1); | |
| 6535 | 61 | qed "program_equalityI"; | 
| 62 | ||
| 63 | val [major,minor] = | |
| 64 | Goal "[| F = G; [| Init F = Init G; Acts F = Acts G |] ==> P |] ==> P"; | |
| 65 | by (rtac minor 1); | |
| 66 | by (auto_tac (claset(), simpset() addsimps [major])); | |
| 67 | qed "program_equalityE"; | |
| 68 | ||
| 8327 | 69 | Goal "(F=G) = (Init F = Init G & Acts F = Acts G)"; | 
| 70 | by (blast_tac (claset() addIs [program_equalityI, program_equalityE]) 1); | |
| 71 | qed "program_equality_iff"; | |
| 72 | ||
| 7915 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7826diff
changeset | 73 | Addsimps [surjective_mk_program]; | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7826diff
changeset | 74 | |
| 6535 | 75 | |
| 76 | (*** These rules allow "lazy" definition expansion | |
| 77 | They avoid expanding the full program, which is a large expression | |
| 78 | ***) | |
| 79 | ||
| 80 | Goal "F == mk_program (init,acts) ==> Init F = init"; | |
| 81 | by Auto_tac; | |
| 82 | qed "def_prg_Init"; | |
| 83 | ||
| 84 | (*The program is not expanded, but its Init and Acts are*) | |
| 85 | val [rew] = goal thy | |
| 86 | "[| F == mk_program (init,acts) |] \ | |
| 87 | \ ==> Init F = init & Acts F = insert Id acts"; | |
| 88 | by (rewtac rew); | |
| 89 | by Auto_tac; | |
| 90 | qed "def_prg_simps"; | |
| 91 | ||
| 92 | (*An action is expanded only if a pair of states is being tested against it*) | |
| 93 | Goal "[| act == {(s,s'). P s s'} |] ==> ((s,s') : act) = P s s'";
 | |
| 94 | by Auto_tac; | |
| 95 | qed "def_act_simp"; | |
| 96 | ||
| 97 | fun simp_of_act def = def RS def_act_simp; | |
| 98 | ||
| 99 | (*A set is expanded only if an element is being tested against it*) | |
| 100 | Goal "A == B ==> (x : A) = (x : B)"; | |
| 101 | by Auto_tac; | |
| 102 | qed "def_set_simp"; | |
| 103 | ||
| 104 | fun simp_of_set def = def RS def_set_simp; | |
| 105 | ||
| 106 | ||
| 6536 | 107 | (*** co ***) | 
| 4776 | 108 | |
| 7403 | 109 | (*These operators are not overloaded, but their operands are sets, and | 
| 110 | ultimately there's a risk of reaching equality, which IS overloaded*) | |
| 111 | overload_1st_set "UNITY.constrains"; | |
| 5648 | 112 | overload_1st_set "UNITY.stable"; | 
| 113 | overload_1st_set "UNITY.unless"; | |
| 5340 | 114 | |
| 5277 
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
 paulson parents: 
5253diff
changeset | 115 | val prems = Goalw [constrains_def] | 
| 5648 | 116 | "(!!act s s'. [| act: Acts F; (s,s') : act; s: A |] ==> s': A') \ | 
| 6536 | 117 | \ ==> F : A co A'"; | 
| 4776 | 118 | by (blast_tac (claset() addIs prems) 1); | 
| 119 | qed "constrainsI"; | |
| 120 | ||
| 5069 | 121 | Goalw [constrains_def] | 
| 6536 | 122 | "[| F : A co A'; act: Acts F; (s,s'): act; s: A |] ==> s': A'"; | 
| 4776 | 123 | by (Blast_tac 1); | 
| 124 | qed "constrainsD"; | |
| 125 | ||
| 6536 | 126 | Goalw [constrains_def] "F : {} co B";
 | 
| 4776 | 127 | by (Blast_tac 1); | 
| 128 | qed "constrains_empty"; | |
| 129 | ||
| 7826 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 130 | Goalw [constrains_def] "(F : A co {}) = (A={})";
 | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 131 | by (Blast_tac 1); | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 132 | qed "constrains_empty2"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 133 | |
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 134 | Goalw [constrains_def] "(F : UNIV co B) = (B = UNIV)"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 135 | by (Blast_tac 1); | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 136 | qed "constrains_UNIV"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 137 | |
| 6536 | 138 | Goalw [constrains_def] "F : A co UNIV"; | 
| 4776 | 139 | by (Blast_tac 1); | 
| 7826 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 140 | qed "constrains_UNIV2"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 141 | |
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 142 | AddIffs [constrains_empty, constrains_empty2, | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 143 | constrains_UNIV, constrains_UNIV2]; | 
| 4776 | 144 | |
| 5648 | 145 | (*monotonic in 2nd argument*) | 
| 5069 | 146 | Goalw [constrains_def] | 
| 6536 | 147 | "[| F : A co A'; A'<=B' |] ==> F : A co B'"; | 
| 4776 | 148 | by (Blast_tac 1); | 
| 149 | qed "constrains_weaken_R"; | |
| 150 | ||
| 5648 | 151 | (*anti-monotonic in 1st argument*) | 
| 5069 | 152 | Goalw [constrains_def] | 
| 6536 | 153 | "[| F : A co A'; B<=A |] ==> F : B co A'"; | 
| 4776 | 154 | by (Blast_tac 1); | 
| 155 | qed "constrains_weaken_L"; | |
| 156 | ||
| 5069 | 157 | Goalw [constrains_def] | 
| 6536 | 158 | "[| F : A co A'; B<=A; A'<=B' |] ==> F : B co B'"; | 
| 4776 | 159 | by (Blast_tac 1); | 
| 160 | qed "constrains_weaken"; | |
| 161 | ||
| 162 | (** Union **) | |
| 163 | ||
| 5069 | 164 | Goalw [constrains_def] | 
| 7345 | 165 | "[| F : A co A'; F : B co B' |] ==> F : (A Un B) co (A' Un B')"; | 
| 4776 | 166 | by (Blast_tac 1); | 
| 167 | qed "constrains_Un"; | |
| 168 | ||
| 5069 | 169 | Goalw [constrains_def] | 
| 7345 | 170 | "ALL i:I. F : (A i) co (A' i) ==> F : (UN i:I. A i) co (UN i:I. A' i)"; | 
| 4776 | 171 | by (Blast_tac 1); | 
| 8334 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8327diff
changeset | 172 | bind_thm ("constrains_UN", ballI RS result());
 | 
| 4776 | 173 | |
| 7826 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 174 | Goalw [constrains_def] "(A Un B) co C = (A co C) Int (B co C)"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 175 | by (Blast_tac 1); | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 176 | qed "constrains_Un_distrib"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 177 | |
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 178 | Goalw [constrains_def] "(UN i:I. A i) co B = (INT i:I. A i co B)"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 179 | by (Blast_tac 1); | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 180 | qed "constrains_UN_distrib"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 181 | |
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 182 | Goalw [constrains_def] "C co (A Int B) = (C co A) Int (C co B)"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 183 | by (Blast_tac 1); | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 184 | qed "constrains_Int_distrib"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 185 | |
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 186 | Goalw [constrains_def] "A co (INT i:I. B i) = (INT i:I. A co B i)"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 187 | by (Blast_tac 1); | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 188 | qed "constrains_INT_distrib"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7630diff
changeset | 189 | |
| 4776 | 190 | (** Intersection **) | 
| 191 | ||
| 5069 | 192 | Goalw [constrains_def] | 
| 7345 | 193 | "[| F : A co A'; F : B co B' |] ==> F : (A Int B) co (A' Int B')"; | 
| 4776 | 194 | by (Blast_tac 1); | 
| 195 | qed "constrains_Int"; | |
| 196 | ||
| 5069 | 197 | Goalw [constrains_def] | 
| 7345 | 198 | "ALL i:I. F : (A i) co (A' i) ==> F : (INT i:I. A i) co (INT i:I. A' i)"; | 
| 4776 | 199 | by (Blast_tac 1); | 
| 8334 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8327diff
changeset | 200 | bind_thm ("constrains_INT", ballI RS result());
 | 
| 4776 | 201 | |
| 6536 | 202 | Goalw [constrains_def] "F : A co A' ==> A <= A'"; | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6012diff
changeset | 203 | by Auto_tac; | 
| 5277 
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
 paulson parents: 
5253diff
changeset | 204 | qed "constrains_imp_subset"; | 
| 4776 | 205 | |
| 6536 | 206 | (*The reasoning is by subsets since "co" refers to single actions | 
| 6012 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5804diff
changeset | 207 | only. So this rule isn't that useful.*) | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6012diff
changeset | 208 | Goalw [constrains_def] | 
| 6536 | 209 | "[| F : A co B; F : B co C |] ==> F : A co C"; | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6012diff
changeset | 210 | by (Blast_tac 1); | 
| 5277 
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
 paulson parents: 
5253diff
changeset | 211 | qed "constrains_trans"; | 
| 4776 | 212 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6012diff
changeset | 213 | Goalw [constrains_def] | 
| 7345 | 214 | "[| F : A co (A' Un B); F : B co B' |] ==> F : A co (A' Un B')"; | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6012diff
changeset | 215 | by (Clarify_tac 1); | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6012diff
changeset | 216 | by (Blast_tac 1); | 
| 6012 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5804diff
changeset | 217 | qed "constrains_cancel"; | 
| 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5804diff
changeset | 218 | |
| 4776 | 219 | |
| 7630 | 220 | (*** unless ***) | 
| 221 | ||
| 222 | Goalw [unless_def] "F : (A-B) co (A Un B) ==> F : A unless B"; | |
| 223 | by (assume_tac 1); | |
| 224 | qed "unlessI"; | |
| 225 | ||
| 226 | Goalw [unless_def] "F : A unless B ==> F : (A-B) co (A Un B)"; | |
| 227 | by (assume_tac 1); | |
| 228 | qed "unlessD"; | |
| 229 | ||
| 230 | ||
| 4776 | 231 | (*** stable ***) | 
| 232 | ||
| 6536 | 233 | Goalw [stable_def] "F : A co A ==> F : stable A"; | 
| 4776 | 234 | by (assume_tac 1); | 
| 235 | qed "stableI"; | |
| 236 | ||
| 6536 | 237 | Goalw [stable_def] "F : stable A ==> F : A co A"; | 
| 4776 | 238 | by (assume_tac 1); | 
| 239 | qed "stableD"; | |
| 240 | ||
| 7594 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7403diff
changeset | 241 | Goalw [stable_def, constrains_def] "stable UNIV = UNIV"; | 
| 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7403diff
changeset | 242 | by Auto_tac; | 
| 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7403diff
changeset | 243 | qed "stable_UNIV"; | 
| 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7403diff
changeset | 244 | Addsimps [stable_UNIV]; | 
| 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7403diff
changeset | 245 | |
| 5804 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5669diff
changeset | 246 | (** Union **) | 
| 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5669diff
changeset | 247 | |
| 5069 | 248 | Goalw [stable_def] | 
| 5648 | 249 | "[| F : stable A; F : stable A' |] ==> F : stable (A Un A')"; | 
| 4776 | 250 | by (blast_tac (claset() addIs [constrains_Un]) 1); | 
| 251 | qed "stable_Un"; | |
| 252 | ||
| 5069 | 253 | Goalw [stable_def] | 
| 5804 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5669diff
changeset | 254 | "ALL i:I. F : stable (A i) ==> F : stable (UN i:I. A i)"; | 
| 8334 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8327diff
changeset | 255 | by (blast_tac (claset() addIs [constrains_UN]) 1); | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8327diff
changeset | 256 | bind_thm ("stable_UN", ballI RS result());
 | 
| 5804 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5669diff
changeset | 257 | |
| 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5669diff
changeset | 258 | (** Intersection **) | 
| 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5669diff
changeset | 259 | |
| 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5669diff
changeset | 260 | Goalw [stable_def] | 
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
7915diff
changeset | 261 | "[| F : stable A; F : stable A' |] ==> F : stable (A Int A')"; | 
| 4776 | 262 | by (blast_tac (claset() addIs [constrains_Int]) 1); | 
| 263 | qed "stable_Int"; | |
| 264 | ||
| 5804 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5669diff
changeset | 265 | Goalw [stable_def] | 
| 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5669diff
changeset | 266 | "ALL i:I. F : stable (A i) ==> F : stable (INT i:I. A i)"; | 
| 8334 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8327diff
changeset | 267 | by (blast_tac (claset() addIs [constrains_INT]) 1); | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8327diff
changeset | 268 | bind_thm ("stable_INT", ballI RS result());
 | 
| 5804 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5669diff
changeset | 269 | |
| 5277 
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
 paulson parents: 
5253diff
changeset | 270 | Goalw [stable_def, constrains_def] | 
| 7345 | 271 | "[| F : stable C; F : A co (C Un A') |] ==> F : (C Un A) co (C Un A')"; | 
| 4776 | 272 | by (Blast_tac 1); | 
| 5277 
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
 paulson parents: 
5253diff
changeset | 273 | qed "stable_constrains_Un"; | 
| 4776 | 274 | |
| 5277 
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
 paulson parents: 
5253diff
changeset | 275 | Goalw [stable_def, constrains_def] | 
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
7915diff
changeset | 276 | "[| F : stable C; F : (C Int A) co A' |] ==> F : (C Int A) co (C Int A')"; | 
| 4776 | 277 | by (Blast_tac 1); | 
| 5277 
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
 paulson parents: 
5253diff
changeset | 278 | qed "stable_constrains_Int"; | 
| 4776 | 279 | |
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
7915diff
changeset | 280 | (*[| F : stable C; F : (C Int A) co A |] ==> F : stable (C Int A) *) | 
| 5648 | 281 | bind_thm ("stable_constrains_stable", stable_constrains_Int RS stableI);
 | 
| 282 | ||
| 283 | ||
| 5804 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5669diff
changeset | 284 | (*** invariant ***) | 
| 5648 | 285 | |
| 286 | Goal "[| Init F<=A; F: stable A |] ==> F : invariant A"; | |
| 287 | by (asm_simp_tac (simpset() addsimps [invariant_def]) 1); | |
| 288 | qed "invariantI"; | |
| 289 | ||
| 290 | (*Could also say "invariant A Int invariant B <= invariant (A Int B)"*) | |
| 291 | Goal "[| F : invariant A; F : invariant B |] ==> F : invariant (A Int B)"; | |
| 292 | by (auto_tac (claset(), simpset() addsimps [invariant_def, stable_Int])); | |
| 293 | qed "invariant_Int"; | |
| 294 | ||
| 295 | ||
| 296 | (*** increasing ***) | |
| 297 | ||
| 8216 | 298 | Goalw [increasing_def] | 
| 299 |      "F : increasing f ==> F : stable {s. z <= f s}";
 | |
| 300 | by (Blast_tac 1); | |
| 301 | qed "increasingD"; | |
| 302 | ||
| 9025 | 303 | Goalw [increasing_def, stable_def] "F : increasing (%s. c)"; | 
| 304 | by Auto_tac; | |
| 305 | qed "increasing_constant"; | |
| 306 | AddIffs [increasing_constant]; | |
| 307 | ||
| 5648 | 308 | Goalw [increasing_def, stable_def, constrains_def] | 
| 6712 | 309 | "mono g ==> increasing f <= increasing (g o f)"; | 
| 5648 | 310 | by Auto_tac; | 
| 6712 | 311 | by (blast_tac (claset() addIs [monoD, order_trans]) 1); | 
| 312 | qed "mono_increasing_o"; | |
| 5648 | 313 | |
| 8216 | 314 | (*Holds by the theorem (Suc m <= n) = (m < n) *) | 
| 5648 | 315 | Goalw [increasing_def] | 
| 8216 | 316 |      "!!z::nat. F : increasing f ==> F: stable {s. z < f s}";
 | 
| 5648 | 317 | by (simp_tac (simpset() addsimps [Suc_le_eq RS sym]) 1); | 
| 318 | by (Blast_tac 1); | |
| 8216 | 319 | qed "strict_increasingD"; | 
| 5648 | 320 | |
| 321 | ||
| 322 | (** The Elimination Theorem. The "free" m has become universally quantified! | |
| 323 | Should the premise be !!m instead of ALL m ? Would make it harder to use | |
| 324 | in forward proof. **) | |
| 325 | ||
| 5069 | 326 | Goalw [constrains_def] | 
| 6536 | 327 |     "[| ALL m:M. F : {s. s x = m} co (B m) |] \
 | 
| 328 | \    ==> F : {s. s x : M} co (UN m:M. B m)";
 | |
| 4776 | 329 | by (Blast_tac 1); | 
| 330 | qed "elimination"; | |
| 331 | ||
| 332 | (*As above, but for the trivial case of a one-variable state, in which the | |
| 333 | state is identified with its one variable.*) | |
| 5069 | 334 | Goalw [constrains_def] | 
| 6536 | 335 |     "(ALL m:M. F : {m} co (B m)) ==> F : M co (UN m:M. B m)";
 | 
| 4776 | 336 | by (Blast_tac 1); | 
| 337 | qed "elimination_sing"; | |
| 338 | ||
| 339 | ||
| 340 | ||
| 341 | (*** Theoretical Results from Section 6 ***) | |
| 342 | ||
| 5069 | 343 | Goalw [constrains_def, strongest_rhs_def] | 
| 6536 | 344 | "F : A co (strongest_rhs F A )"; | 
| 4776 | 345 | by (Blast_tac 1); | 
| 346 | qed "constrains_strongest_rhs"; | |
| 347 | ||
| 5069 | 348 | Goalw [constrains_def, strongest_rhs_def] | 
| 6536 | 349 | "F : A co B ==> strongest_rhs F A <= B"; | 
| 4776 | 350 | by (Blast_tac 1); | 
| 351 | qed "strongest_rhs_is_strongest"; | |
| 8948 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 352 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 353 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 354 | (** Ad-hoc set-theory rules **) | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 355 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 356 | Goal "A Un B - (A - B) = B"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 357 | by (Blast_tac 1); | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 358 | qed "Un_Diff_Diff"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 359 | Addsimps [Un_Diff_Diff]; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 360 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 361 | Goal "Union(B) Int A = Union((%C. C Int A)``B)"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 362 | by (Blast_tac 1); | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 363 | qed "Int_Union_Union"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 364 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 365 | (** Needed for WF reasoning in WFair.ML **) | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 366 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 367 | Goal "less_than ^^ {k} = greaterThan k";
 | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 368 | by (Blast_tac 1); | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 369 | qed "Image_less_than"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 370 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 371 | Goal "less_than^-1 ^^ {k} = lessThan k";
 | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 372 | by (Blast_tac 1); | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 373 | qed "Image_inverse_less_than"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 374 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8703diff
changeset | 375 | Addsimps [Image_less_than, Image_inverse_less_than]; |