| author | wenzelm | 
| Tue, 13 Jun 2000 18:33:34 +0200 | |
| changeset 9058 | 7856a01119fb | 
| parent 9021 | 5643223dad0a | 
| child 13550 | 5a176b8dda84 | 
| permissions | -rw-r--r-- | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 1 | (* Title: HOL/UNITY/Constrains | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 4 | Copyright 1998 University of Cambridge | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 5 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 6 | Safety relations: restricted to the set of reachable states. | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 7 | *) | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 8 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 9 | |
| 6535 | 10 | (*** traces and reachable ***) | 
| 11 | ||
| 12 | Goal "reachable F = {s. EX evs. (s,evs): traces (Init F) (Acts F)}";
 | |
| 13 | by Safe_tac; | |
| 14 | by (etac traces.induct 2); | |
| 15 | by (etac reachable.induct 1); | |
| 16 | by (ALLGOALS (blast_tac (claset() addIs reachable.intrs @ traces.intrs))); | |
| 17 | qed "reachable_equiv_traces"; | |
| 18 | ||
| 19 | Goal "Init F <= reachable F"; | |
| 20 | by (blast_tac (claset() addIs reachable.intrs) 1); | |
| 21 | qed "Init_subset_reachable"; | |
| 22 | ||
| 23 | Goal "Acts G <= Acts F ==> G : stable (reachable F)"; | |
| 24 | by (blast_tac (claset() addIs [stableI, constrainsI] @ reachable.intrs) 1); | |
| 25 | qed "stable_reachable"; | |
| 26 | ||
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
7689diff
changeset | 27 | AddSIs [stable_reachable]; | 
| 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
7689diff
changeset | 28 | Addsimps [stable_reachable]; | 
| 6535 | 29 | |
| 30 | (*The set of all reachable states is an invariant...*) | |
| 31 | Goal "F : invariant (reachable F)"; | |
| 32 | by (simp_tac (simpset() addsimps [invariant_def]) 1); | |
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
7689diff
changeset | 33 | by (blast_tac (claset() addIs reachable.intrs) 1); | 
| 6535 | 34 | qed "invariant_reachable"; | 
| 35 | ||
| 36 | (*...in fact the strongest invariant!*) | |
| 37 | Goal "F : invariant A ==> reachable F <= A"; | |
| 38 | by (full_simp_tac | |
| 39 | (simpset() addsimps [stable_def, constrains_def, invariant_def]) 1); | |
| 40 | by (rtac subsetI 1); | |
| 41 | by (etac reachable.induct 1); | |
| 42 | by (REPEAT (blast_tac (claset() addIs reachable.intrs) 1)); | |
| 43 | qed "invariant_includes_reachable"; | |
| 44 | ||
| 45 | ||
| 6536 | 46 | (*** Co ***) | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 47 | |
| 7403 | 48 | (*Needed because its operands are sets*) | 
| 49 | overload_1st_set "Constrains.Constrains"; | |
| 5340 | 50 | |
| 6536 | 51 | (*F : B co B' ==> F : (reachable F Int B) co (reachable F Int B')*) | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 52 | bind_thm ("constrains_reachable_Int",
 | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 53 | subset_refl RS | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 54 | rewrite_rule [stable_def] stable_reachable RS | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 55 | constrains_Int); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 56 | |
| 6575 | 57 | (*Resembles the previous definition of Constrains*) | 
| 58 | Goalw [Constrains_def] | |
| 59 |      "A Co B = {F. F : (reachable F  Int  A) co (reachable F  Int  B)}";
 | |
| 60 | by (blast_tac (claset() addDs [constrains_reachable_Int] | |
| 61 | addIs [constrains_weaken]) 1); | |
| 62 | qed "Constrains_eq_constrains"; | |
| 63 | ||
| 6536 | 64 | Goalw [Constrains_def] "F : A co A' ==> F : A Co A'"; | 
| 6575 | 65 | by (blast_tac (claset() addIs [constrains_weaken_L]) 1); | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 66 | qed "constrains_imp_Constrains"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 67 | |
| 5648 | 68 | Goalw [stable_def, Stable_def] "F : stable A ==> F : Stable A"; | 
| 5631 | 69 | by (etac constrains_imp_Constrains 1); | 
| 70 | qed "stable_imp_Stable"; | |
| 71 | ||
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 72 | val prems = Goal | 
| 5620 | 73 | "(!!act s s'. [| act: Acts F; (s,s') : act; s: A |] ==> s': A') \ | 
| 6536 | 74 | \ ==> F : A Co A'"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 75 | by (rtac constrains_imp_Constrains 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 76 | by (blast_tac (claset() addIs (constrainsI::prems)) 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 77 | qed "ConstrainsI"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 78 | |
| 6536 | 79 | Goalw [Constrains_def, constrains_def] "F : {} Co B";
 | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 80 | by (Blast_tac 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 81 | qed "Constrains_empty"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 82 | |
| 6536 | 83 | Goal "F : A Co UNIV"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 84 | by (blast_tac (claset() addIs [ConstrainsI]) 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 85 | qed "Constrains_UNIV"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 86 | AddIffs [Constrains_empty, Constrains_UNIV]; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 87 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 88 | Goalw [Constrains_def] | 
| 6536 | 89 | "[| F : A Co A'; A'<=B' |] ==> F : A Co B'"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 90 | by (blast_tac (claset() addIs [constrains_weaken_R]) 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 91 | qed "Constrains_weaken_R"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 92 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 93 | Goalw [Constrains_def] | 
| 6536 | 94 | "[| F : A Co A'; B<=A |] ==> F : B Co A'"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 95 | by (blast_tac (claset() addIs [constrains_weaken_L]) 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 96 | qed "Constrains_weaken_L"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 97 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 98 | Goalw [Constrains_def] | 
| 6536 | 99 | "[| F : A Co A'; B<=A; A'<=B' |] ==> F : B Co B'"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 100 | by (blast_tac (claset() addIs [constrains_weaken]) 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 101 | qed "Constrains_weaken"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 102 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 103 | (** Union **) | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 104 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 105 | Goalw [Constrains_def] | 
| 6536 | 106 | "[| F : A Co A'; F : B Co B' |] \ | 
| 107 | \ ==> F : (A Un B) Co (A' Un B')"; | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 108 | by (blast_tac (claset() addIs [constrains_Un RS constrains_weaken]) 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 109 | qed "Constrains_Un"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 110 | |
| 8334 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 111 | val [prem] = Goalw [Constrains_def] | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 112 | "(!!i. i:I ==> F : (A i) Co (A' i)) \ | 
| 6536 | 113 | \ ==> F : (UN i:I. A i) Co (UN i:I. A' i)"; | 
| 8334 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 114 | by (rtac CollectI 1); | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 115 | by (rtac (prem RS CollectD RS constrains_UN RS constrains_weaken) 1); | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 116 | by Auto_tac; | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 117 | qed "Constrains_UN"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 118 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 119 | (** Intersection **) | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 120 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 121 | Goalw [Constrains_def] | 
| 6536 | 122 | "[| F : A Co A'; F : B Co B' |] \ | 
| 123 | \ ==> F : (A Int B) Co (A' Int B')"; | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 124 | by (blast_tac (claset() addIs [constrains_Int RS constrains_weaken]) 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 125 | qed "Constrains_Int"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 126 | |
| 8334 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 127 | val [prem] = Goalw [Constrains_def] | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 128 | "(!!i. i:I ==> F : (A i) Co (A' i)) \ | 
| 6536 | 129 | \ ==> F : (INT i:I. A i) Co (INT i:I. A' i)"; | 
| 8334 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 130 | by (rtac CollectI 1); | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 131 | by (rtac (prem RS CollectD RS constrains_INT RS constrains_weaken) 1); | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 132 | by Auto_tac; | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 133 | qed "Constrains_INT"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 134 | |
| 6536 | 135 | Goal "F : A Co A' ==> reachable F Int A <= A'"; | 
| 6575 | 136 | by (asm_full_simp_tac (simpset() addsimps [constrains_imp_subset, | 
| 137 | Constrains_def]) 1); | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 138 | qed "Constrains_imp_subset"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 139 | |
| 6575 | 140 | Goal "[| F : A Co B; F : B Co C |] ==> F : A Co C"; | 
| 141 | by (full_simp_tac (simpset() addsimps [Constrains_eq_constrains]) 1); | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6012diff
changeset | 142 | by (blast_tac (claset() addIs [constrains_trans, constrains_weaken]) 1); | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 143 | qed "Constrains_trans"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 144 | |
| 6575 | 145 | Goal "[| F : A Co (A' Un B); F : B Co B' |] ==> F : A Co (A' Un B')"; | 
| 146 | by (full_simp_tac (simpset() addsimps [Constrains_eq_constrains, | |
| 147 | constrains_def]) 1); | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6012diff
changeset | 148 | by (Blast_tac 1); | 
| 6012 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5804diff
changeset | 149 | qed "Constrains_cancel"; | 
| 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5804diff
changeset | 150 | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 151 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 152 | (*** Stable ***) | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 153 | |
| 8334 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 154 | (*Useful because there's no Stable_weaken. [Tanja Vos]*) | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 155 | Goal "[| F: Stable A; A = B |] ==> F : Stable B"; | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 156 | by (Blast_tac 1); | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 157 | qed "Stable_eq"; | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 158 | |
| 5648 | 159 | Goal "(F : Stable A) = (F : stable (reachable F Int A))"; | 
| 6575 | 160 | by (simp_tac (simpset() addsimps [Stable_def, Constrains_eq_constrains, | 
| 161 | stable_def]) 1); | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 162 | qed "Stable_eq_stable"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 163 | |
| 6536 | 164 | Goalw [Stable_def] "F : A Co A ==> F : Stable A"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 165 | by (assume_tac 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 166 | qed "StableI"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 167 | |
| 6536 | 168 | Goalw [Stable_def] "F : Stable A ==> F : A Co A"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 169 | by (assume_tac 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 170 | qed "StableD"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 171 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 172 | Goalw [Stable_def] | 
| 5648 | 173 | "[| F : Stable A; F : Stable A' |] ==> F : Stable (A Un A')"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 174 | by (blast_tac (claset() addIs [Constrains_Un]) 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 175 | qed "Stable_Un"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 176 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 177 | Goalw [Stable_def] | 
| 5648 | 178 | "[| F : Stable A; F : Stable A' |] ==> F : Stable (A Int A')"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 179 | by (blast_tac (claset() addIs [Constrains_Int]) 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 180 | qed "Stable_Int"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 181 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 182 | Goalw [Stable_def] | 
| 6536 | 183 | "[| F : Stable C; F : A Co (C Un A') |] \ | 
| 184 | \ ==> F : (C Un A) Co (C Un A')"; | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 185 | by (blast_tac (claset() addIs [Constrains_Un RS Constrains_weaken]) 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 186 | qed "Stable_Constrains_Un"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 187 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 188 | Goalw [Stable_def] | 
| 6536 | 189 | "[| F : Stable C; F : (C Int A) Co A' |] \ | 
| 190 | \ ==> F : (C Int A) Co (C Int A')"; | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 191 | by (blast_tac (claset() addIs [Constrains_Int RS Constrains_weaken]) 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 192 | qed "Stable_Constrains_Int"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 193 | |
| 8334 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 194 | val [prem] = Goalw [Stable_def] | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 195 | "(!!i. i:I ==> F : Stable (A i)) ==> F : Stable (UN i:I. A i)"; | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 196 | by (rtac (prem RS Constrains_UN) 1); | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 197 | by (assume_tac 1); | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 198 | qed "Stable_UN"; | 
| 5804 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5784diff
changeset | 199 | |
| 8334 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 200 | val [prem] = Goalw [Stable_def] | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 201 | "(!!i. i:I ==> F : Stable (A i)) ==> F : Stable (INT i:I. A i)"; | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 202 | by (rtac (prem RS Constrains_INT) 1); | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 203 | by (assume_tac 1); | 
| 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 204 | qed "Stable_INT"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 205 | |
| 5648 | 206 | Goal "F : Stable (reachable F)"; | 
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
7689diff
changeset | 207 | by (simp_tac (simpset() addsimps [Stable_eq_stable]) 1); | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 208 | qed "Stable_reachable"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 209 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 210 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 211 | |
| 5784 | 212 | (*** Increasing ***) | 
| 213 | ||
| 8216 | 214 | Goalw [Increasing_def] | 
| 215 |      "F : Increasing f ==> F : Stable {s. x <= f s}";
 | |
| 216 | by (Blast_tac 1); | |
| 217 | qed "IncreasingD"; | |
| 218 | ||
| 5784 | 219 | Goalw [Increasing_def, Stable_def, Constrains_def, stable_def, constrains_def] | 
| 6704 | 220 | "mono g ==> Increasing f <= Increasing (g o f)"; | 
| 5784 | 221 | by Auto_tac; | 
| 6704 | 222 | by (blast_tac (claset() addIs [monoD, order_trans]) 1); | 
| 223 | qed "mono_Increasing_o"; | |
| 5784 | 224 | |
| 225 | Goalw [Increasing_def] | |
| 8216 | 226 |      "!!z::nat. F : Increasing f ==> F: Stable {s. z < f s}";
 | 
| 5784 | 227 | by (simp_tac (simpset() addsimps [Suc_le_eq RS sym]) 1); | 
| 228 | by (Blast_tac 1); | |
| 8216 | 229 | qed "strict_IncreasingD"; | 
| 5784 | 230 | |
| 231 | Goalw [increasing_def, Increasing_def] | |
| 232 | "F : increasing f ==> F : Increasing f"; | |
| 233 | by (blast_tac (claset() addIs [stable_imp_Stable]) 1); | |
| 234 | qed "increasing_imp_Increasing"; | |
| 235 | ||
| 9021 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 236 | bind_thm ("Increasing_constant", 
 | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 237 | increasing_constant RS increasing_imp_Increasing); | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 238 | AddIffs [Increasing_constant]; | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 239 | |
| 5784 | 240 | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 241 | (*** The Elimination Theorem. The "free" m has become universally quantified! | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 242 | Should the premise be !!m instead of ALL m ? Would make it harder to use | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 243 | in forward proof. ***) | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 244 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 245 | Goalw [Constrains_def, constrains_def] | 
| 6536 | 246 |     "[| ALL m. F : {s. s x = m} Co (B m) |] \
 | 
| 247 | \    ==> F : {s. s x : M} Co (UN m:M. B m)";
 | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 248 | by (Blast_tac 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 249 | qed "Elimination"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 250 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 251 | (*As above, but for the trivial case of a one-variable state, in which the | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 252 | state is identified with its one variable.*) | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 253 | Goalw [Constrains_def, constrains_def] | 
| 6536 | 254 |     "(ALL m. F : {m} Co (B m)) ==> F : M Co (UN m:M. B m)";
 | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 255 | by (Blast_tac 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 256 | qed "Elimination_sing"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 257 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 258 | |
| 6570 | 259 | (*** Specialized laws for handling Always ***) | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 260 | |
| 6570 | 261 | (** Natural deduction rules for "Always A" **) | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 262 | |
| 6570 | 263 | Goal "[| Init F<=A; F : Stable A |] ==> F : Always A"; | 
| 264 | by (asm_simp_tac (simpset() addsimps [Always_def]) 1); | |
| 265 | qed "AlwaysI"; | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 266 | |
| 6570 | 267 | Goal "F : Always A ==> Init F<=A & F : Stable A"; | 
| 268 | by (asm_full_simp_tac (simpset() addsimps [Always_def]) 1); | |
| 269 | qed "AlwaysD"; | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 270 | |
| 6570 | 271 | bind_thm ("AlwaysE", AlwaysD RS conjE);
 | 
| 8334 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
8216diff
changeset | 272 | bind_thm ("Always_imp_Stable", AlwaysD RS conjunct2);
 | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 273 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 274 | |
| 6570 | 275 | (*The set of all reachable states is Always*) | 
| 276 | Goal "F : Always A ==> reachable F <= A"; | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 277 | by (full_simp_tac | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 278 | (simpset() addsimps [Stable_def, Constrains_def, constrains_def, | 
| 6570 | 279 | Always_def]) 1); | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 280 | by (rtac subsetI 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 281 | by (etac reachable.induct 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 282 | by (REPEAT (blast_tac (claset() addIs reachable.intrs) 1)); | 
| 6570 | 283 | qed "Always_includes_reachable"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 284 | |
| 6575 | 285 | Goalw [Always_def, invariant_def, Stable_def, stable_def] | 
| 6570 | 286 | "F : invariant A ==> F : Always A"; | 
| 6575 | 287 | by (blast_tac (claset() addIs [constrains_imp_Constrains]) 1); | 
| 6570 | 288 | qed "invariant_imp_Always"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 289 | |
| 6672 | 290 | bind_thm ("Always_reachable", invariant_reachable RS invariant_imp_Always);
 | 
| 291 | ||
| 6575 | 292 | Goal "Always A = {F. F : invariant (reachable F Int A)}";
 | 
| 293 | by (simp_tac (simpset() addsimps [Always_def, invariant_def, Stable_def, | |
| 294 | Constrains_eq_constrains, stable_def]) 1); | |
| 5648 | 295 | by (blast_tac (claset() addIs reachable.intrs) 1); | 
| 6570 | 296 | qed "Always_eq_invariant_reachable"; | 
| 5648 | 297 | |
| 6570 | 298 | (*the RHS is the traditional definition of the "always" operator*) | 
| 299 | Goal "Always A = {F. reachable F <= A}";
 | |
| 5804 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5784diff
changeset | 300 | by (auto_tac (claset() addDs [invariant_includes_reachable], | 
| 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5784diff
changeset | 301 | simpset() addsimps [Int_absorb2, invariant_reachable, | 
| 6570 | 302 | Always_eq_invariant_reachable])); | 
| 303 | qed "Always_eq_includes_reachable"; | |
| 5648 | 304 | |
| 7689 | 305 | Goal "Always UNIV = UNIV"; | 
| 306 | by (auto_tac (claset(), | |
| 307 | simpset() addsimps [Always_eq_includes_reachable])); | |
| 308 | qed "Always_UNIV_eq"; | |
| 309 | Addsimps [Always_UNIV_eq]; | |
| 5648 | 310 | |
| 9021 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 311 | Goal "UNIV <= A ==> F : Always A"; | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 312 | by (auto_tac (claset(), simpset() addsimps [Always_eq_includes_reachable])); | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 313 | qed "UNIV_AlwaysI"; | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 314 | |
| 6570 | 315 | Goal "Always A = (UN I: Pow A. invariant I)"; | 
| 316 | by (simp_tac (simpset() addsimps [Always_eq_includes_reachable]) 1); | |
| 5804 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5784diff
changeset | 317 | by (blast_tac (claset() addIs [invariantI, impOfSubs Init_subset_reachable, | 
| 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5784diff
changeset | 318 | impOfSubs invariant_includes_reachable]) 1); | 
| 6570 | 319 | qed "Always_eq_UN_invariant"; | 
| 320 | ||
| 321 | Goal "[| F : Always A; A <= B |] ==> F : Always B"; | |
| 322 | by (auto_tac (claset(), simpset() addsimps [Always_eq_includes_reachable])); | |
| 323 | qed "Always_weaken"; | |
| 5804 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5784diff
changeset | 324 | |
| 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5784diff
changeset | 325 | |
| 6570 | 326 | (*** "Co" rules involving Always ***) | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 327 | |
| 6739 
66e4118eead9
replaced rules Always_ConstrainsI/D by equivalences Always_Constrains_pre,
 paulson parents: 
6704diff
changeset | 328 | Goal "F : Always INV ==> (F : (INV Int A) Co A') = (F : A Co A')"; | 
| 
66e4118eead9
replaced rules Always_ConstrainsI/D by equivalences Always_Constrains_pre,
 paulson parents: 
6704diff
changeset | 329 | by (asm_simp_tac | 
| 6570 | 330 | (simpset() addsimps [Always_includes_reachable RS Int_absorb2, | 
| 5804 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 paulson parents: 
5784diff
changeset | 331 | Constrains_def, Int_assoc RS sym]) 1); | 
| 6739 
66e4118eead9
replaced rules Always_ConstrainsI/D by equivalences Always_Constrains_pre,
 paulson parents: 
6704diff
changeset | 332 | qed "Always_Constrains_pre"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 333 | |
| 6739 
66e4118eead9
replaced rules Always_ConstrainsI/D by equivalences Always_Constrains_pre,
 paulson parents: 
6704diff
changeset | 334 | Goal "F : Always INV ==> (F : A Co (INV Int A')) = (F : A Co A')"; | 
| 
66e4118eead9
replaced rules Always_ConstrainsI/D by equivalences Always_Constrains_pre,
 paulson parents: 
6704diff
changeset | 335 | by (asm_simp_tac | 
| 6570 | 336 | (simpset() addsimps [Always_includes_reachable RS Int_absorb2, | 
| 6575 | 337 | Constrains_eq_constrains, Int_assoc RS sym]) 1); | 
| 6739 
66e4118eead9
replaced rules Always_ConstrainsI/D by equivalences Always_Constrains_pre,
 paulson parents: 
6704diff
changeset | 338 | qed "Always_Constrains_post"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 339 | |
| 6739 
66e4118eead9
replaced rules Always_ConstrainsI/D by equivalences Always_Constrains_pre,
 paulson parents: 
6704diff
changeset | 340 | (* [| F : Always INV; F : (INV Int A) Co A' |] ==> F : A Co A' *) | 
| 
66e4118eead9
replaced rules Always_ConstrainsI/D by equivalences Always_Constrains_pre,
 paulson parents: 
6704diff
changeset | 341 | bind_thm ("Always_ConstrainsI", Always_Constrains_pre RS iffD1);
 | 
| 
66e4118eead9
replaced rules Always_ConstrainsI/D by equivalences Always_Constrains_pre,
 paulson parents: 
6704diff
changeset | 342 | |
| 
66e4118eead9
replaced rules Always_ConstrainsI/D by equivalences Always_Constrains_pre,
 paulson parents: 
6704diff
changeset | 343 | (* [| F : Always INV; F : A Co A' |] ==> F : A Co (INV Int A') *) | 
| 
66e4118eead9
replaced rules Always_ConstrainsI/D by equivalences Always_Constrains_pre,
 paulson parents: 
6704diff
changeset | 344 | bind_thm ("Always_ConstrainsD", Always_Constrains_post RS iffD2);
 | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 345 | |
| 9021 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 346 | (*The analogous proof of Always_LeadsTo_weaken doesn't terminate*) | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 347 | Goal "[| F : Always C; F : A Co A'; \ | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 348 | \ C Int B <= A; C Int A' <= B' |] \ | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 349 | \ ==> F : B Co B'"; | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 350 | by (rtac Always_ConstrainsI 1); | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 351 | by (assume_tac 1); | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 352 | by (dtac Always_ConstrainsD 1); | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 353 | by (assume_tac 1); | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 354 | by (blast_tac (claset() addIs [Constrains_weaken]) 1); | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 355 | qed "Always_Constrains_weaken"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 356 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 357 | |
| 6570 | 358 | (** Conjoining Always properties **) | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 359 | |
| 7541 
1a7a38d8f5bd
new theorem Always_INT_distrib; therefore renamed Always_Int
 paulson parents: 
7403diff
changeset | 360 | Goal "Always (A Int B) = Always A Int Always B"; | 
| 
1a7a38d8f5bd
new theorem Always_INT_distrib; therefore renamed Always_Int
 paulson parents: 
7403diff
changeset | 361 | by (auto_tac (claset(), simpset() addsimps [Always_eq_includes_reachable])); | 
| 
1a7a38d8f5bd
new theorem Always_INT_distrib; therefore renamed Always_Int
 paulson parents: 
7403diff
changeset | 362 | qed "Always_Int_distrib"; | 
| 
1a7a38d8f5bd
new theorem Always_INT_distrib; therefore renamed Always_Int
 paulson parents: 
7403diff
changeset | 363 | |
| 
1a7a38d8f5bd
new theorem Always_INT_distrib; therefore renamed Always_Int
 paulson parents: 
7403diff
changeset | 364 | Goal "Always (INTER I A) = (INT i:I. Always (A i))"; | 
| 
1a7a38d8f5bd
new theorem Always_INT_distrib; therefore renamed Always_Int
 paulson parents: 
7403diff
changeset | 365 | by (auto_tac (claset(), simpset() addsimps [Always_eq_includes_reachable])); | 
| 
1a7a38d8f5bd
new theorem Always_INT_distrib; therefore renamed Always_Int
 paulson parents: 
7403diff
changeset | 366 | qed "Always_INT_distrib"; | 
| 
1a7a38d8f5bd
new theorem Always_INT_distrib; therefore renamed Always_Int
 paulson parents: 
7403diff
changeset | 367 | |
| 6570 | 368 | Goal "[| F : Always A; F : Always B |] ==> F : Always (A Int B)"; | 
| 9021 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 369 | by (asm_simp_tac (simpset() addsimps [Always_Int_distrib]) 1); | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 370 | qed "Always_Int_I"; | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 371 | |
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 372 | (*Allows a kind of "implication introduction"*) | 
| 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 373 | Goal "F : Always A ==> (F : Always (-A Un B)) = (F : Always B)"; | 
| 6570 | 374 | by (auto_tac (claset(), simpset() addsimps [Always_eq_includes_reachable])); | 
| 9021 
5643223dad0a
new theorems Always_Constrains_weaken and Always_Compl_Un_eq
 paulson parents: 
8334diff
changeset | 375 | qed "Always_Compl_Un_eq"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 376 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 377 | (*Delete the nearest invariance assumption (which will be the second one | 
| 7541 
1a7a38d8f5bd
new theorem Always_INT_distrib; therefore renamed Always_Int
 paulson parents: 
7403diff
changeset | 378 | used by Always_Int_I) *) | 
| 6570 | 379 | val Always_thin = | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 380 | read_instantiate_sg (sign_of thy) | 
| 6570 | 381 |                 [("V", "?F : Always ?A")] thin_rl;
 | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 382 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 383 | (*Combines two invariance ASSUMPTIONS into one. USEFUL??*) | 
| 7541 
1a7a38d8f5bd
new theorem Always_INT_distrib; therefore renamed Always_Int
 paulson parents: 
7403diff
changeset | 384 | val Always_Int_tac = dtac Always_Int_I THEN' assume_tac THEN' etac Always_thin; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 385 | |
| 5319 
7356d0c88b1b
Moved Un_subset_iff and Int_subset_iff from UNITY to equalities.ML
 paulson parents: 
5313diff
changeset | 386 | (*Combines a list of invariance THEOREMS into one.*) | 
| 7541 
1a7a38d8f5bd
new theorem Always_INT_distrib; therefore renamed Always_Int
 paulson parents: 
7403diff
changeset | 387 | val Always_Int_rule = foldr1 (fn (th1,th2) => [th1,th2] MRS Always_Int_I); | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 388 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 389 | |
| 5648 | 390 | (*To allow expansion of the program's definition when appropriate*) | 
| 391 | val program_defs_ref = ref ([] : thm list); | |
| 392 | ||
| 6536 | 393 | (*proves "co" properties when the program is specified*) | 
| 5648 | 394 | fun constrains_tac i = | 
| 5422 | 395 | SELECT_GOAL | 
| 7403 | 396 | (EVERY [REPEAT (Always_Int_tac 1), | 
| 397 | REPEAT (etac Always_ConstrainsI 1 | |
| 398 | ORELSE | |
| 399 | resolve_tac [StableI, stableI, | |
| 5422 | 400 | constrains_imp_Constrains] 1), | 
| 401 | rtac constrainsI 1, | |
| 7403 | 402 | full_simp_tac (simpset() addsimps !program_defs_ref) 1, | 
| 5620 | 403 | REPEAT (FIRSTGOAL (etac disjE)), | 
| 5422 | 404 | ALLGOALS Clarify_tac, | 
| 5648 | 405 | ALLGOALS Asm_full_simp_tac]) i; | 
| 7403 | 406 | |
| 407 | ||
| 408 | (*For proving invariants*) | |
| 409 | fun always_tac i = | |
| 410 | rtac AlwaysI i THEN Force_tac i THEN constrains_tac i; |