author | wenzelm |
Wed, 03 Feb 1999 16:42:40 +0100 | |
changeset 6188 | c40e5ac04e3e |
parent 6019 | 0e55c2fb2ebb |
child 6295 | 351b3c2b0d83 |
permissions | -rw-r--r-- |
5252 | 1 |
(* Title: HOL/UNITY/Union.ML |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
5 |
||
6 |
Unions of programs |
|
7 |
||
8 |
From Misra's Chapter 5: Asynchronous Compositions of Programs |
|
9 |
*) |
|
10 |
||
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
11 |
Goal "k:I ==> (INT i:I. A i) Int A k = (INT i:I. A i)"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
12 |
by (Blast_tac 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
13 |
qed "INT_absorb2"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
14 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
15 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
16 |
val rinst = read_instantiate_sg (sign_of thy); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
17 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
18 |
Addcongs [UN_cong, INT_cong]; |
5252 | 19 |
|
5867 | 20 |
(** SKIP **) |
21 |
||
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
22 |
Goal "States (SKIP A) = A"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
23 |
by (simp_tac (simpset() addsimps [SKIP_def]) 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
24 |
qed "States_SKIP"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
25 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
26 |
Goal "Init (SKIP A) = A"; |
5648 | 27 |
by (simp_tac (simpset() addsimps [SKIP_def]) 1); |
5611 | 28 |
qed "Init_SKIP"; |
29 |
||
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
30 |
Goal "Acts (SKIP A) = {diag A}"; |
5648 | 31 |
by (simp_tac (simpset() addsimps [SKIP_def]) 1); |
32 |
qed "Acts_SKIP"; |
|
33 |
||
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
34 |
Addsimps [States_SKIP, Init_SKIP, Acts_SKIP]; |
5648 | 35 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
36 |
Goal "reachable (SKIP A) = A"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
37 |
by (force_tac (claset() addEs [reachable.induct] |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
38 |
addIs reachable.intrs, simpset()) 1); |
5867 | 39 |
qed "reachable_SKIP"; |
40 |
||
41 |
Addsimps [reachable_SKIP]; |
|
42 |
||
43 |
||
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
44 |
(** Join **) |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
45 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
46 |
Goal "States (F Join G) = States F Int States G"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
47 |
by (simp_tac (simpset() addsimps [Join_def]) 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
48 |
qed "States_Join"; |
5867 | 49 |
|
5611 | 50 |
Goal "Init (F Join G) = Init F Int Init G"; |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
51 |
by (auto_tac (claset() addIs [impOfSubs Init_subset_States], |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
52 |
simpset() addsimps [Join_def])); |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
53 |
qed "Init_Join"; |
5252 | 54 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
55 |
Goal "States F = States G ==> Acts (F Join G) = Acts F Un Acts G"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
56 |
by (subgoal_tac "Acts F Un Acts G <= Pow (States G Times States G)" 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
57 |
by (blast_tac (claset() addEs [equalityE] |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
58 |
addDs [impOfSubs Acts_subset_Pow_States]) 2); |
5596 | 59 |
by (auto_tac (claset(), simpset() addsimps [Join_def])); |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
60 |
qed "Acts_Join"; |
5252 | 61 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
62 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
63 |
(** JN **) |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
64 |
|
6019 | 65 |
Goalw [JOIN_def, SKIP_def] "(JN i:{}. F i) = SKIP UNIV"; |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
66 |
by Auto_tac; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
67 |
qed "JN_empty"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
68 |
Addsimps [JN_empty]; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
69 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
70 |
Goal "States (JN i:I. F i) = (INT i:I. States (F i))"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
71 |
by (simp_tac (simpset() addsimps [JOIN_def]) 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
72 |
qed "States_JN"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
73 |
|
5611 | 74 |
Goal "Init (JN i:I. F i) = (INT i:I. Init (F i))"; |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
75 |
by (auto_tac (claset() addIs [impOfSubs Init_subset_States], |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
76 |
simpset() addsimps [JOIN_def])); |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
77 |
qed "Init_JN"; |
5252 | 78 |
|
6019 | 79 |
(*If I={} then the LHS is (SKIP UNIV) while the RHS is {}.*) |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
80 |
Goalw [eqStates_def] |
6019 | 81 |
"[| eqStates I F; i: I |] ==> Acts (JN i:I. F i) = (UN i:I. Acts (F i))"; |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
82 |
by (Clarify_tac 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
83 |
by (subgoal_tac "(UN i:I. Acts (F i)) <= Pow (St Times St)" 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
84 |
by (blast_tac (claset() addEs [equalityE] |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
85 |
addDs [impOfSubs Acts_subset_Pow_States]) 2); |
5596 | 86 |
by (auto_tac (claset(), simpset() addsimps [JOIN_def])); |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
87 |
qed "Acts_JN"; |
5252 | 88 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
89 |
Goal "eqStates I F \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
90 |
\ ==> Acts (JN i:I. F i) = \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
91 |
\ (if I={} then {diag UNIV} else (UN i:I. Acts (F i)))"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
92 |
by (force_tac (claset(), simpset() addsimps [Acts_JN]) 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
93 |
qed "Acts_JN_if"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
94 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
95 |
Addsimps [States_Join, Init_Join, States_JN, Init_JN]; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
96 |
|
5259 | 97 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
98 |
Goal "(JN x:insert j I. F x) = (F j) Join (JN x:I. F x)"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
99 |
by (rtac program_equalityI 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
100 |
by (ALLGOALS Asm_simp_tac); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
101 |
by (asm_simp_tac |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
102 |
(simpset() addsimps [JOIN_def, Join_def, Acts_eq_raw, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
103 |
image_UNION, image_Un, image_image, Int_assoc]) 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
104 |
qed "JN_insert"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
105 |
Addsimps[JN_insert]; |
5867 | 106 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
107 |
Goal "k:I ==> A k Join (JN i:I. A i) = (JN i:I. A i)"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
108 |
by (stac (JN_insert RS sym) 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
109 |
by (asm_simp_tac (simpset() addsimps [insert_absorb]) 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
110 |
qed "JN_absorb"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
111 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
112 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
113 |
Goalw [eqStates_def] "eqStates {} F"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
114 |
by (Simp_tac 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
115 |
qed "eqStates_empty"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
116 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
117 |
Goalw [eqStates_def] |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
118 |
"[| eqStates I F; States (F i) = (INT i: I. States (F i)) |] \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
119 |
\ ==> eqStates (insert i I) F"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
120 |
by Auto_tac; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
121 |
qed "eqStates_insertI"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
122 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
123 |
Goalw [eqStates_def] |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
124 |
"eqStates (insert i I) F = \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
125 |
\ (eqStates I F & (I={} | States (F i) = (INT i: I. States (F i))))"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
126 |
by Auto_tac; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
127 |
qed "eqStates_insert_eq"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
128 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
129 |
Addsimps [eqStates_empty, eqStates_insert_eq]; |
5584 | 130 |
|
131 |
||
5611 | 132 |
(** Algebraic laws **) |
5259 | 133 |
|
5611 | 134 |
Goal "F Join G = G Join F"; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
135 |
by (simp_tac (simpset() addsimps [Join_def, Un_commute, Int_commute]) 1); |
5259 | 136 |
qed "Join_commute"; |
137 |
||
5611 | 138 |
Goal "(F Join G) Join H = F Join (G Join H)"; |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
139 |
by (rtac program_equalityI 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
140 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps Un_ac@[Int_assoc]))); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
141 |
by (asm_simp_tac |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
142 |
(simpset() addsimps Un_ac@Int_ac@[Join_def, Acts_eq_raw, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
143 |
image_Un, image_image]) 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
144 |
by (Blast_tac 1); |
5259 | 145 |
qed "Join_assoc"; |
5596 | 146 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
147 |
Goalw [Join_def, SKIP_def] "States F <= A ==> (SKIP A) Join F = F"; |
5611 | 148 |
by (rtac program_equalityI 1); |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
149 |
by (ALLGOALS |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
150 |
(asm_simp_tac (simpset() addsimps |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
151 |
Int_ac@[Acts_subset_Pow_States RS restrict_rel_image, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
152 |
Acts_eq_raw, insert_absorb, Int_absorb1]))); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
153 |
by (blast_tac (claset() addIs [impOfSubs Init_subset_States]) 1); |
5611 | 154 |
qed "Join_SKIP_left"; |
5584 | 155 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
156 |
Goal "States F <= A ==> F Join (SKIP A) = F"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
157 |
by (stac Join_commute 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
158 |
by (asm_simp_tac (simpset() addsimps [Join_SKIP_left]) 1); |
5611 | 159 |
qed "Join_SKIP_right"; |
160 |
||
161 |
Addsimps [Join_SKIP_left, Join_SKIP_right]; |
|
162 |
||
163 |
Goalw [Join_def] "F Join F = F"; |
|
164 |
by (rtac program_equalityI 1); |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
165 |
by (ALLGOALS |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
166 |
(asm_simp_tac (simpset() addsimps |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
167 |
[insert_absorb, Acts_subset_Pow_States RS Acts_eq]))); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
168 |
by (blast_tac (claset() addIs [impOfSubs Init_subset_States]) 1); |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
169 |
qed "Join_absorb"; |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
170 |
|
5611 | 171 |
Addsimps [Join_absorb]; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
172 |
|
5970 | 173 |
Goal "(JN i: I Un J. A i) = ((JN i: I. A i) Join (JN i:J. A i))"; |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
174 |
by (rtac program_equalityI 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
175 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [INT_Un]))); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
176 |
by (asm_simp_tac |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
177 |
(simpset() addsimps Int_ac@ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
178 |
[JOIN_def, Join_def, Acts_eq_raw, UN_Un, INT_Un, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
179 |
image_UNION, image_Un, image_image]) 1); |
5970 | 180 |
qed "JN_Join"; |
181 |
||
182 |
Goal "i: I ==> (JN i:I. c) = c"; |
|
183 |
by (auto_tac (claset() addSIs [program_equalityI], |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
184 |
simpset() addsimps [Acts_JN, eqStates_def])); |
5970 | 185 |
qed "JN_constant"; |
186 |
||
187 |
Goal "(JN i:I. A i Join B i) = (JN i:I. A i) Join (JN i:I. B i)"; |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
188 |
by (rtac program_equalityI 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
189 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [INT_Int_distrib]))); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
190 |
by (asm_simp_tac |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
191 |
(simpset() addsimps [JOIN_def, Join_def, Acts_eq_raw, image_UNION, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
192 |
rinst [("A","%x. States (A x) Int States (B x)")] |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
193 |
INT_absorb2, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
194 |
image_Un, image_image]) 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
195 |
by (asm_simp_tac (simpset() addsimps [INT_Int_distrib RS sym] @ Int_ac) 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
196 |
by (Blast_tac 1); |
5970 | 197 |
qed "JN_Join_distrib"; |
5584 | 198 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
199 |
Goal "i : I ==> (JN i:I. A i Join B) = ((JN i:I. A i) Join B)"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
200 |
by (asm_simp_tac (simpset() addsimps [JN_Join_distrib, JN_constant]) 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
201 |
qed "JN_Join_miniscope"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
202 |
|
5584 | 203 |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
204 |
(*** Safety: constrains, stable, FP ***) |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
205 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
206 |
Goal "[| F : constrains A B; G : constrains A B |] \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
207 |
\ ==> F Join G : constrains A B"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
208 |
by (auto_tac (claset(), |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
209 |
simpset() addsimps [constrains_def, Join_def, Acts_eq_raw, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
210 |
image_Un])); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
211 |
by (REPEAT (force_tac (claset() addSEs [ballE], simpset()) 1)); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
212 |
qed "constrains_imp_Join_constrains"; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
213 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
214 |
Goal "States F = States G ==> \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
215 |
\ F Join G : constrains A B = (F : constrains A B & G : constrains A B)"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
216 |
by (auto_tac (claset(), simpset() addsimps [constrains_def, Acts_Join])); |
5620 | 217 |
qed "constrains_Join"; |
218 |
||
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
219 |
Goal "[| i : I; ALL i:I. F i : constrains A B |] \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
220 |
\ ==> (JN i:I. F i) : constrains A B"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
221 |
by (full_simp_tac (simpset() addsimps [constrains_def, JOIN_def, Acts_eq_raw, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
222 |
image_Un]) 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
223 |
by Safe_tac; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
224 |
by (REPEAT (force_tac (claset() addSEs [ballE], simpset()) 1)); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
225 |
qed "constrains_imp_JN_constrains"; |
5620 | 226 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
227 |
Goal "[| i : I; eqStates I F |] \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
228 |
\ ==> (JN i:I. F i) : constrains A B = (ALL i:I. F i : constrains A B)"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
229 |
by (auto_tac (claset(), simpset() addsimps [constrains_def, Acts_JN])); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
230 |
qed "constrains_JN"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
231 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
232 |
(**FAILS, I think; see 5.4.1, Substitution Axiom. |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
233 |
The problem is to relate reachable (F Join G) with |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
234 |
reachable F and reachable G |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
235 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
236 |
Goalw [Constrains_def] |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
237 |
"(JN i:I. F i) : Constrains A B = (ALL i:I. F i : Constrains A B)"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
238 |
by (simp_tac (simpset() addsimps [constrains_JN]) 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
239 |
by (Blast_tac 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
240 |
qed "Constrains_JN"; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
241 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
242 |
Goal "F Join G : Constrains A B = (Constrains F A B & Constrains G A B)"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
243 |
by (auto_tac |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
244 |
(claset(), |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
245 |
simpset() addsimps [Constrains_def, constrains_Join])); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
246 |
qed "Constrains_Join"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
247 |
**) |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
248 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
249 |
Goal "[| F : constrains A A'; G : constrains B B'; States F = States G |] \ |
5648 | 250 |
\ ==> F Join G : constrains (A Int B) (A' Un B')"; |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
251 |
by (asm_simp_tac (simpset() addsimps [constrains_Join]) 1); |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
252 |
by (blast_tac (claset() addIs [constrains_weaken]) 1); |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
253 |
qed "constrains_Join_weaken"; |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
254 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
255 |
Goal "[| i : I; eqStates I F |] \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
256 |
\ ==> (JN i:I. F i) : stable A = (ALL i:I. F i : stable A)"; |
5584 | 257 |
by (asm_simp_tac (simpset() addsimps [stable_def, constrains_JN]) 1); |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
258 |
qed "stable_JN"; |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
259 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
260 |
Goal "[| ALL i:I. F i : invariant A; i : I; eqStates I F |] \ |
5970 | 261 |
\ ==> (JN i:I. F i) : invariant A"; |
262 |
by (asm_full_simp_tac (simpset() addsimps [invariant_def, stable_JN]) 1); |
|
263 |
by (Blast_tac 1); |
|
264 |
bind_thm ("invariant_JN_I", ballI RS result()); |
|
265 |
||
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
266 |
Goal "States F = States G \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
267 |
\ ==> F Join G : stable A = (F : stable A & G : stable A)"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
268 |
by (asm_simp_tac (simpset() addsimps [stable_def, constrains_Join]) 1); |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
269 |
qed "stable_Join"; |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
270 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
271 |
Goal "[| F : invariant A; G : invariant A; States F = States G |] \ |
5970 | 272 |
\ ==> F Join G : invariant A"; |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
273 |
by (asm_full_simp_tac (simpset() addsimps [invariant_def, stable_Join]) 1); |
5970 | 274 |
by (Blast_tac 1); |
275 |
qed "invariant_JoinI"; |
|
276 |
||
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
277 |
Goal "[| i : I; eqStates I F |] \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
278 |
\ ==> FP (JN i:I. F i) = (INT i:I. FP (F i))"; |
5584 | 279 |
by (asm_simp_tac (simpset() addsimps [FP_def, stable_JN, INTER_def]) 1); |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
280 |
qed "FP_JN"; |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
281 |
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
282 |
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
283 |
(*** Progress: transient, ensures ***) |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
284 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
285 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
286 |
Goal "[| (JN i:I. F i) : transient A; i: I |] ==> EX i:I. F i : transient A"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
287 |
by (full_simp_tac (simpset() addsimps [transient_def, JOIN_def, Acts_eq_raw, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
288 |
Int_absorb1, restrict_rel_def]) 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
289 |
by Safe_tac; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
290 |
by (REPEAT (blast_tac (claset() addSIs [rev_bexI]) 1)); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
291 |
qed "transient_JN_imp_transient"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
292 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
293 |
Goal "[| i : I; eqStates I F |] \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
294 |
\ ==> (JN i:I. F i) : transient A = (EX i:I. F i : transient A)"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
295 |
by (auto_tac (claset() addSIs [transient_JN_imp_transient], simpset())); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
296 |
by (auto_tac (claset(), simpset() addsimps [transient_def, Acts_JN])); |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
297 |
qed "transient_JN"; |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
298 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
299 |
Goal "F Join G : transient A ==> \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
300 |
\ F : transient A | G : transient A"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
301 |
by (full_simp_tac (simpset() addsimps [transient_def, Join_def, Acts_eq_raw, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
302 |
restrict_rel_def]) 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
303 |
by Safe_tac; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
304 |
(*delete act:Acts F possibility*) |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
305 |
by (rtac FalseE 3); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
306 |
(*delete act:Acts G possibility*) |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
307 |
by (thin_tac "~ (EX act:Acts G. ?P act)" 2); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
308 |
by (REPEAT (blast_tac (claset() addSIs [rev_bexI]) 1)); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
309 |
qed "transient_Join_imp_transient"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
310 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
311 |
Goal "States F = States G \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
312 |
\ ==> (F Join G : transient A) = (F : transient A | G : transient A)"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
313 |
by (auto_tac (claset() addSIs [transient_Join_imp_transient], simpset())); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
314 |
by (auto_tac (claset(), simpset() addsimps [transient_def, Acts_Join])); |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
315 |
qed "transient_Join"; |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
316 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
317 |
Goal "[| i : I; eqStates I F |] \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
318 |
\ ==> (JN i:I. F i) : ensures A B = \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
319 |
\ ((ALL i:I. F i : constrains (A-B) (A Un B)) & \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
320 |
\ (EX i:I. F i : ensures A B))"; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
321 |
by (auto_tac (claset(), |
5584 | 322 |
simpset() addsimps [ensures_def, constrains_JN, transient_JN])); |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
323 |
qed "ensures_JN"; |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
324 |
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
325 |
Goalw [ensures_def] |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
326 |
"States F = States G \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
327 |
\ ==> F Join G : ensures A B = \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
328 |
\ (F : constrains (A-B) (A Un B) & \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
329 |
\ G : constrains (A-B) (A Un B) & \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
330 |
\ (F : ensures A B | G : ensures A B))"; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
331 |
by (auto_tac (claset(), |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
332 |
simpset() addsimps [constrains_Join, transient_Join])); |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
333 |
qed "ensures_Join"; |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
334 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
335 |
Goal "[| F : stable A; G : constrains A A'; \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
336 |
\ States F = States G; A <= States G |] \ |
5648 | 337 |
\ ==> F Join G : constrains A A'"; |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
338 |
by (forward_tac [constrains_imp_subset] 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
339 |
by (assume_tac 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
340 |
by (asm_full_simp_tac (simpset() addsimps [stable_def, constrains_def, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
341 |
ball_Un, Acts_Join]) 1); |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
342 |
by (Blast_tac 1); |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
343 |
qed "stable_constrains_Join"; |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
344 |
|
5648 | 345 |
(*Premise for G cannot use Invariant because Stable F A is weaker than |
346 |
G : stable A *) |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
347 |
Goal "[| F : stable A; G : invariant A; \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
348 |
\ States F = States G |] ==> F Join G : Invariant A"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
349 |
by (asm_full_simp_tac (simpset() addsimps [Invariant_def, invariant_def, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
350 |
Stable_eq_stable, stable_Join]) 1); |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
351 |
by (force_tac(claset() addIs [stable_reachable, stable_Int], |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
352 |
simpset() addsimps [Acts_Join]) 1); |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
353 |
qed "stable_Join_Invariant"; |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
354 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
355 |
Goal "[| F : stable A; G : ensures A B; \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
356 |
\ States F = States G |] ==> F Join G : ensures A B"; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
357 |
by (asm_simp_tac (simpset() addsimps [ensures_Join]) 1); |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
358 |
by (asm_full_simp_tac (simpset() addsimps [stable_def, ensures_def]) 1); |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
359 |
by (etac constrains_weaken 1); |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
360 |
by Auto_tac; |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
361 |
qed "ensures_stable_Join1"; |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
362 |
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
363 |
(*As above, but exchanging the roles of F and G*) |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
364 |
Goal "[| F : ensures A B; G : stable A; \ |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
365 |
\ States F = States G |] ==> F Join G : ensures A B"; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
366 |
by (stac Join_commute 1); |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
367 |
by (dtac sym 1); |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
368 |
by (blast_tac (claset() addIs [ensures_stable_Join1]) 1); |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
369 |
qed "ensures_stable_Join2"; |
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
370 |
|
5648 | 371 |
|
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
372 |
(** Diff and localTo **) |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
373 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
374 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
375 |
Goal "States (Diff F acts) = States F"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
376 |
by (simp_tac (simpset() addsimps [Diff_def]) 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
377 |
qed "States_Diff"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
378 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
379 |
Goal "Init (Diff F acts) = Init F"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
380 |
by (auto_tac (claset() addIs [impOfSubs Init_subset_States], |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
381 |
simpset() addsimps [Diff_def])); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
382 |
qed "Init_Diff"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
383 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
384 |
Goal "Acts (Diff F acts) = insert (diag (States F)) (Acts F - acts)"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
385 |
by (subgoal_tac "Acts F - acts <= Pow (States F Times States F)" 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
386 |
by (blast_tac (claset() addEs [equalityE] |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
387 |
addDs [impOfSubs Acts_subset_Pow_States]) 2); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
388 |
by (auto_tac (claset(), simpset() addsimps [Diff_def])); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
389 |
qed "Acts_Diff"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
390 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
391 |
Addsimps [States_Diff, Init_Diff, Acts_Diff]; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
392 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
393 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
394 |
Goalw [Join_def] "States F = States G ==> F Join (Diff G (Acts F)) = F Join G"; |
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
395 |
by (rtac program_equalityI 1); |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
396 |
by (auto_tac (claset(), simpset() addsimps [insert_absorb])); |
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
397 |
qed "Join_Diff2"; |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
398 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
399 |
Goalw [Disjoint_def] "States F = States G ==> Disjoint F (Diff G (Acts F))"; |
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
400 |
by Auto_tac; |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
401 |
qed "Diff_Disjoint"; |
5648 | 402 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
403 |
Goalw [Disjoint_def] "Disjoint F G ==> States F = States G"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
404 |
by Auto_tac; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
405 |
qed "Disjoint_States_eq"; |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
406 |
|
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
407 |
Goal "[| F Join G : v localTo F; Disjoint F G |] \ |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
408 |
\ ==> G : (INT z. stable {s. v s = z})"; |
5648 | 409 |
by (asm_full_simp_tac |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
410 |
(simpset() addsimps [localTo_def, Disjoint_def, |
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
411 |
Acts_Join, stable_def, constrains_def]) 1); |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
412 |
by (Blast_tac 1); |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
413 |
qed "localTo_imp_stable"; |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
414 |
|
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
415 |
Goal "[| F Join G : v localTo F; (s,s') : act; \ |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
416 |
\ act : Acts G; Disjoint F G |] ==> v s' = v s"; |
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
417 |
by (asm_full_simp_tac |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
418 |
(simpset() addsimps [localTo_def, Disjoint_def, |
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
419 |
Acts_Join, stable_def, constrains_def]) 1); |
5648 | 420 |
by (Blast_tac 1); |
421 |
qed "localTo_imp_equals"; |
|
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
422 |
|
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
423 |
Goalw [localTo_def, stable_def, constrains_def] |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
424 |
"v localTo F <= (f o v) localTo F"; |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
425 |
by (Clarify_tac 1); |
5898 | 426 |
by (force_tac (claset() addSEs [allE, ballE], simpset()) 1); |
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
427 |
qed "localTo_imp_o_localTo"; |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
428 |
|
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
429 |
|
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
430 |
(*** Higher-level rules involving localTo and Join ***) |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
431 |
|
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
432 |
Goal "[| F : constrains {s. P (v s) (w s)} {s. P' (v s) (w s)}; \ |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
433 |
\ F Join G : v localTo F; \ |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
434 |
\ F Join G : w localTo F; \ |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
435 |
\ Disjoint F G |] \ |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
436 |
\ ==> F Join G : constrains {s. P (v s) (w s)} {s. P' (v s) (w s)}"; |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
437 |
by (auto_tac (claset(), |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
438 |
simpset() addsimps [constrains_def, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
439 |
Disjoint_States_eq RS Acts_Join])); |
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
440 |
by (REPEAT_FIRST (dtac localTo_imp_equals THEN' REPEAT1 o atac)); |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
441 |
by (subgoal_tac "xa : States F" 1); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
442 |
by (force_tac |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
443 |
(claset() addSDs [Disjoint_States_eq,impOfSubs Acts_subset_Pow_States], |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
444 |
simpset()) 2); |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
445 |
by (Force_tac 1); |
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
446 |
qed "constrains_localTo_constrains2"; |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
447 |
|
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
448 |
Goalw [stable_def] |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
449 |
"[| F : stable {s. P (v s) (w s)}; \ |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
450 |
\ F Join G : v localTo F; \ |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
451 |
\ F Join G : w localTo F; \ |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
452 |
\ Disjoint F G |] \ |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
453 |
\ ==> F Join G : stable {s. P (v s) (w s)}"; |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
454 |
by (blast_tac (claset() addIs [constrains_localTo_constrains2]) 1); |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
455 |
qed "stable_localTo_stable2"; |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
456 |
|
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
457 |
|
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
458 |
Goal "(UN k. {s. f s = k}) = UNIV"; |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
459 |
by (Blast_tac 1); |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
460 |
qed "UN_eq_UNIV"; |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
461 |
|
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
462 |
Goal "[| F : stable {s. v s <= w s}; \ |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
463 |
\ F Join G : v localTo F; \ |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
464 |
\ F Join G : Increasing w; \ |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
465 |
\ Disjoint F G |] \ |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
466 |
\ ==> F Join G : Stable {s. v s <= w s}"; |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
467 |
by (safe_tac (claset() addSDs [localTo_imp_stable])); |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
468 |
by (rewrite_goals_tac [stable_def, Stable_def, Increasing_def]); |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
469 |
by (subgoal_tac "ALL k: UNIV. ?H : Constrains ({s. v s = k} Int ?AA) ?AA" 1); |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
470 |
by (dtac ball_Constrains_UN 1); |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
471 |
by (full_simp_tac (simpset() addsimps [UN_eq_UNIV]) 1); |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
472 |
by (rtac ballI 1); |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
473 |
by (subgoal_tac "F Join G : constrains ({s. v s = k} Int {s. v s <= w s}) \ |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
474 |
\ ({s. v s = k} Un {s. v s <= w s})" 1); |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
475 |
by (asm_simp_tac |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
476 |
(simpset() addsimps [Disjoint_States_eq RS constrains_Join]) 2); |
5804
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
477 |
by (blast_tac (claset() addIs [constrains_weaken]) 2); |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
478 |
by (dtac (constrains_imp_Constrains RS Constrains_Int) 1 THEN etac INT_D 1); |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
479 |
by (etac Constrains_weaken 2); |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
480 |
by Auto_tac; |
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents:
5785
diff
changeset
|
481 |
qed "Increasing_localTo_Stable"; |