| author | wenzelm |
| Thu, 31 Aug 2000 00:15:09 +0200 | |
| changeset 9760 | 72c0a12ae3bf |
| parent 9687 | 772ac061bd76 |
| child 10064 | 1a77667b21ef |
| 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 |
||
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
11 |
|
| 5867 | 12 |
(** SKIP **) |
13 |
||
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
14 |
Goal "Init SKIP = UNIV"; |
| 5648 | 15 |
by (simp_tac (simpset() addsimps [SKIP_def]) 1); |
| 5611 | 16 |
qed "Init_SKIP"; |
17 |
||
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
18 |
Goal "Acts SKIP = {Id}";
|
| 5648 | 19 |
by (simp_tac (simpset() addsimps [SKIP_def]) 1); |
20 |
qed "Acts_SKIP"; |
|
21 |
||
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
22 |
Addsimps [Init_SKIP, Acts_SKIP]; |
| 5648 | 23 |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
24 |
Goal "reachable SKIP = UNIV"; |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
25 |
by (force_tac (claset() addEs [reachable.induct] |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
26 |
addIs reachable.intrs, simpset()) 1); |
| 5867 | 27 |
qed "reachable_SKIP"; |
28 |
||
29 |
Addsimps [reachable_SKIP]; |
|
30 |
||
| 6836 | 31 |
(** SKIP and safety properties **) |
32 |
||
33 |
Goalw [constrains_def] "(SKIP : A co B) = (A<=B)"; |
|
34 |
by Auto_tac; |
|
35 |
qed "SKIP_in_constrains_iff"; |
|
36 |
AddIffs [SKIP_in_constrains_iff]; |
|
37 |
||
38 |
Goalw [Constrains_def] "(SKIP : A Co B) = (A<=B)"; |
|
39 |
by Auto_tac; |
|
40 |
qed "SKIP_in_Constrains_iff"; |
|
41 |
AddIffs [SKIP_in_Constrains_iff]; |
|
42 |
||
43 |
Goalw [stable_def] "SKIP : stable A"; |
|
44 |
by Auto_tac; |
|
45 |
qed "SKIP_in_stable"; |
|
46 |
AddIffs [SKIP_in_stable, SKIP_in_stable RS stable_imp_Stable]; |
|
47 |
||
| 5867 | 48 |
|
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
49 |
(** Join **) |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
50 |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
51 |
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
|
52 |
by (simp_tac (simpset() addsimps [Join_def]) 1); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
53 |
qed "Init_Join"; |
| 5252 | 54 |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
55 |
Goal "Acts (F Join G) = Acts F Un Acts G"; |
| 5596 | 56 |
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
|
57 |
qed "Acts_Join"; |
| 5252 | 58 |
|
| 7537 | 59 |
Addsimps [Init_Join, Acts_Join]; |
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
60 |
|
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
61 |
|
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
62 |
(** JN **) |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
63 |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
64 |
Goalw [JOIN_def, SKIP_def] "(JN i:{}. F i) = SKIP";
|
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
65 |
by Auto_tac; |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
66 |
qed "JN_empty"; |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
67 |
Addsimps [JN_empty]; |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
68 |
|
| 6633 | 69 |
Goal "(JN i:insert a I. F i) = (F a) Join (JN i:I. F i)"; |
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
70 |
by (rtac program_equalityI 1); |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
71 |
by (ALLGOALS (simp_tac (simpset() addsimps [JOIN_def, Join_def]))); |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
72 |
qed "JN_insert"; |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
73 |
Addsimps[JN_empty, JN_insert]; |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
74 |
|
| 5611 | 75 |
Goal "Init (JN i:I. F i) = (INT i:I. Init (F i))"; |
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
76 |
by (simp_tac (simpset() addsimps [JOIN_def]) 1); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
77 |
qed "Init_JN"; |
| 5252 | 78 |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
79 |
Goal "Acts (JN i:I. F i) = insert Id (UN i:I. Acts (F i))"; |
| 5596 | 80 |
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
|
81 |
qed "Acts_JN"; |
| 5252 | 82 |
|
| 7537 | 83 |
Addsimps [Init_JN, Acts_JN]; |
| 5867 | 84 |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
85 |
val prems = Goalw [JOIN_def] |
| 6633 | 86 |
"[| I=J; !!i. i:J ==> F i = G i |] ==> \ |
87 |
\ (JN i:I. F i) = (JN i:J. G i)"; |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
88 |
by (asm_simp_tac (simpset() addsimps prems) 1); |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
89 |
qed "JN_cong"; |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
90 |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
91 |
Addcongs [JN_cong]; |
| 5584 | 92 |
|
93 |
||
| 5611 | 94 |
(** Algebraic laws **) |
| 5259 | 95 |
|
| 5611 | 96 |
Goal "F Join G = G Join F"; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
97 |
by (simp_tac (simpset() addsimps [Join_def, Un_commute, Int_commute]) 1); |
| 5259 | 98 |
qed "Join_commute"; |
99 |
||
| 7360 | 100 |
Goal "A Join (B Join C) = B Join (A Join C)"; |
101 |
by (simp_tac (simpset() addsimps Un_ac@Int_ac@[Join_def]) 1); |
|
102 |
qed "Join_left_commute"; |
|
103 |
||
| 5611 | 104 |
Goal "(F Join G) Join H = F Join (G Join H)"; |
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
105 |
by (simp_tac (simpset() addsimps Un_ac@[Join_def, Int_assoc]) 1); |
| 5259 | 106 |
qed "Join_assoc"; |
| 5596 | 107 |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
108 |
Goalw [Join_def, SKIP_def] "SKIP Join F = F"; |
| 5611 | 109 |
by (rtac program_equalityI 1); |
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
110 |
by (ALLGOALS (simp_tac (simpset() addsimps [insert_absorb]))); |
| 5611 | 111 |
qed "Join_SKIP_left"; |
| 5584 | 112 |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
113 |
Goalw [Join_def, SKIP_def] "F Join SKIP = F"; |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
114 |
by (rtac program_equalityI 1); |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
115 |
by (ALLGOALS (simp_tac (simpset() addsimps [insert_absorb]))); |
| 5611 | 116 |
qed "Join_SKIP_right"; |
117 |
||
118 |
Addsimps [Join_SKIP_left, Join_SKIP_right]; |
|
119 |
||
120 |
Goalw [Join_def] "F Join F = F"; |
|
121 |
by (rtac program_equalityI 1); |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
122 |
by Auto_tac; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
123 |
qed "Join_absorb"; |
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
124 |
|
| 5611 | 125 |
Addsimps [Join_absorb]; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
126 |
|
|
7915
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
paulson
parents:
7878
diff
changeset
|
127 |
Goalw [Join_def] "F Join (F Join G) = F Join G"; |
| 7360 | 128 |
by (rtac program_equalityI 1); |
129 |
by Auto_tac; |
|
130 |
qed "Join_left_absorb"; |
|
131 |
||
132 |
(*Join is an AC-operator*) |
|
133 |
val Join_ac = [Join_assoc, Join_left_absorb, Join_commute, Join_left_commute]; |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
134 |
|
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
135 |
|
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
136 |
(*** JN laws ***) |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
137 |
|
| 6633 | 138 |
(*Also follows by JN_insert and insert_absorb, but the proof is longer*) |
139 |
Goal "k:I ==> F k Join (JN i:I. F i) = (JN i:I. F i)"; |
|
| 7537 | 140 |
by (auto_tac (claset() addSIs [program_equalityI], simpset())); |
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
141 |
qed "JN_absorb"; |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
142 |
|
| 6633 | 143 |
Goal "(JN i: I Un J. F i) = ((JN i: I. F i) Join (JN i:J. F i))"; |
| 7537 | 144 |
by (auto_tac (claset() addSIs [program_equalityI], simpset())); |
| 6633 | 145 |
qed "JN_Un"; |
| 5970 | 146 |
|
| 6836 | 147 |
Goal "(JN i:I. c) = (if I={} then SKIP else c)";
|
| 8314 | 148 |
by (rtac program_equalityI 1); |
149 |
by Auto_tac; |
|
| 5970 | 150 |
qed "JN_constant"; |
151 |
||
| 6633 | 152 |
Goal "(JN i:I. F i Join G i) = (JN i:I. F i) Join (JN i:I. G i)"; |
| 7537 | 153 |
by (auto_tac (claset() addSIs [program_equalityI], simpset())); |
| 5970 | 154 |
qed "JN_Join_distrib"; |
| 5584 | 155 |
|
| 6633 | 156 |
Goal "i : I ==> (JN i:I. F i Join G) = ((JN i:I. F i) Join G)"; |
157 |
by (asm_simp_tac (simpset() addsimps [JN_Join_distrib, JN_constant]) 1); |
|
| 6836 | 158 |
by Auto_tac; |
| 6633 | 159 |
qed "JN_Join_miniscope"; |
160 |
||
|
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8042
diff
changeset
|
161 |
(*Used to prove guarantees_JN_I*) |
|
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8042
diff
changeset
|
162 |
Goalw [JOIN_def, Join_def] "i: I ==> F i Join JOIN (I - {i}) F = JOIN I F";
|
|
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8042
diff
changeset
|
163 |
by (rtac program_equalityI 1); |
|
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8042
diff
changeset
|
164 |
by Auto_tac; |
|
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8042
diff
changeset
|
165 |
qed "JN_Join_diff"; |
|
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8042
diff
changeset
|
166 |
|
| 5584 | 167 |
|
| 6536 | 168 |
(*** Safety: co, stable, FP ***) |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
169 |
|
|
8251
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8216
diff
changeset
|
170 |
(*Fails if I={} because it collapses to SKIP : A co B, i.e. to A<=B. So an
|
|
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8216
diff
changeset
|
171 |
alternative precondition is A<=B, but most proofs using this rule require |
|
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8216
diff
changeset
|
172 |
I to be nonempty for other reasons anyway.*) |
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
173 |
Goalw [constrains_def, JOIN_def] |
| 6633 | 174 |
"i : I ==> (JN i:I. F i) : A co B = (ALL i:I. F i : A co B)"; |
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
175 |
by (Simp_tac 1); |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
176 |
by (Blast_tac 1); |
| 7523 | 177 |
qed "JN_constrains"; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
178 |
|
| 6633 | 179 |
Goal "(F Join G : A co B) = (F : A co B & G : A co B)"; |
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
180 |
by (auto_tac |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
181 |
(claset(), |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
182 |
simpset() addsimps [constrains_def, Join_def])); |
| 7523 | 183 |
qed "Join_constrains"; |
| 5620 | 184 |
|
|
8122
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
185 |
Goal "(F Join G : A unless B) = (F : A unless B & G : A unless B)"; |
|
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
186 |
by (simp_tac (simpset() addsimps [Join_constrains, unless_def]) 1); |
|
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
187 |
qed "Join_unless"; |
|
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8069
diff
changeset
|
188 |
|
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
8314
diff
changeset
|
189 |
Addsimps [Join_constrains, Join_unless]; |
|
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
8314
diff
changeset
|
190 |
|
| 6633 | 191 |
(*Analogous weak versions FAIL; see Misra [1994] 5.4.1, Substitution Axiom. |
192 |
reachable (F Join G) could be much bigger than reachable F, reachable G |
|
193 |
*) |
|
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5970
diff
changeset
|
194 |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
195 |
|
| 6536 | 196 |
Goal "[| F : A co A'; G : B co B' |] \ |
197 |
\ ==> F Join G : (A Int B) co (A' Un B')"; |
|
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
8314
diff
changeset
|
198 |
by (Simp_tac 1); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
199 |
by (blast_tac (claset() addIs [constrains_weaken]) 1); |
| 7523 | 200 |
qed "Join_constrains_weaken"; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
201 |
|
|
8251
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8216
diff
changeset
|
202 |
(*If I={}, it degenerates to SKIP : UNIV co {}, which is false.*)
|
| 6633 | 203 |
Goal "[| ALL i:I. F i : A i co A' i; i: I |] \ |
204 |
\ ==> (JN i:I. F i) : (INT i:I. A i) co (UN i:I. A' i)"; |
|
| 7523 | 205 |
by (asm_simp_tac (simpset() addsimps [JN_constrains]) 1); |
| 6633 | 206 |
by (blast_tac (claset() addIs [constrains_weaken]) 1); |
| 7523 | 207 |
qed "JN_constrains_weaken"; |
| 6633 | 208 |
|
|
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8042
diff
changeset
|
209 |
Goal "(JN i:I. F i) : stable A = (ALL i:I. F i : stable A)"; |
|
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8042
diff
changeset
|
210 |
by (asm_simp_tac |
|
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8042
diff
changeset
|
211 |
(simpset() addsimps [stable_def, constrains_def, JOIN_def]) 1); |
| 7523 | 212 |
qed "JN_stable"; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
213 |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
214 |
Goal "[| ALL i:I. F i : invariant A; i : I |] \ |
| 5970 | 215 |
\ ==> (JN i:I. F i) : invariant A"; |
| 7523 | 216 |
by (asm_full_simp_tac (simpset() addsimps [invariant_def, JN_stable]) 1); |
| 5970 | 217 |
by (Blast_tac 1); |
218 |
bind_thm ("invariant_JN_I", ballI RS result());
|
|
219 |
||
|
7594
8a188ef6545e
working version with co-guarantees-leadsto results
paulson
parents:
7537
diff
changeset
|
220 |
Goal "(F Join G : stable A) = \ |
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
221 |
\ (F : stable A & G : stable A)"; |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
8314
diff
changeset
|
222 |
by (simp_tac (simpset() addsimps [stable_def]) 1); |
| 7523 | 223 |
qed "Join_stable"; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
224 |
|
|
7594
8a188ef6545e
working version with co-guarantees-leadsto results
paulson
parents:
7537
diff
changeset
|
225 |
Goal "(F Join G : increasing f) = \ |
|
8a188ef6545e
working version with co-guarantees-leadsto results
paulson
parents:
7537
diff
changeset
|
226 |
\ (F : increasing f & G : increasing f)"; |
|
8a188ef6545e
working version with co-guarantees-leadsto results
paulson
parents:
7537
diff
changeset
|
227 |
by (simp_tac (simpset() addsimps [increasing_def, Join_stable]) 1); |
|
8a188ef6545e
working version with co-guarantees-leadsto results
paulson
parents:
7537
diff
changeset
|
228 |
by (Blast_tac 1); |
|
8a188ef6545e
working version with co-guarantees-leadsto results
paulson
parents:
7537
diff
changeset
|
229 |
qed "Join_increasing"; |
|
8a188ef6545e
working version with co-guarantees-leadsto results
paulson
parents:
7537
diff
changeset
|
230 |
|
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
8314
diff
changeset
|
231 |
Addsimps [Join_stable, Join_increasing]; |
|
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
8314
diff
changeset
|
232 |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
233 |
Goal "[| F : invariant A; G : invariant A |] \ |
| 5970 | 234 |
\ ==> F Join G : invariant A"; |
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
8314
diff
changeset
|
235 |
by (full_simp_tac (simpset() addsimps [invariant_def]) 1); |
| 5970 | 236 |
by (Blast_tac 1); |
237 |
qed "invariant_JoinI"; |
|
238 |
||
|
8055
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents:
8042
diff
changeset
|
239 |
Goal "FP (JN i:I. F i) = (INT i:I. FP (F i))"; |
| 7523 | 240 |
by (asm_simp_tac (simpset() addsimps [FP_def, JN_stable, INTER_def]) 1); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
241 |
qed "FP_JN"; |
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
242 |
|
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
243 |
|
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
244 |
(*** Progress: transient, ensures ***) |
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
245 |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
246 |
Goal "i : I ==> \ |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
247 |
\ (JN i:I. F i) : transient A = (EX i:I. F i : transient A)"; |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
248 |
by (auto_tac (claset(), |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
249 |
simpset() addsimps [transient_def, JOIN_def])); |
| 7523 | 250 |
qed "JN_transient"; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
251 |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
252 |
Goal "F Join G : transient A = \ |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
253 |
\ (F : transient A | G : transient A)"; |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
254 |
by (auto_tac (claset(), |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
255 |
simpset() addsimps [bex_Un, transient_def, |
|
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
256 |
Join_def])); |
| 7523 | 257 |
qed "Join_transient"; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
258 |
|
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
8314
diff
changeset
|
259 |
Addsimps [Join_transient]; |
|
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
8314
diff
changeset
|
260 |
|
| 8216 | 261 |
Goal "F : transient A ==> F Join G : transient A"; |
262 |
by (asm_simp_tac (simpset() addsimps [Join_transient]) 1); |
|
263 |
qed "Join_transient_I1"; |
|
264 |
||
265 |
Goal "G : transient A ==> F Join G : transient A"; |
|
266 |
by (asm_simp_tac (simpset() addsimps [Join_transient]) 1); |
|
267 |
qed "Join_transient_I2"; |
|
268 |
||
|
8251
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8216
diff
changeset
|
269 |
(*If I={} it degenerates to (SKIP : A ensures B) = False, i.e. to ~(A<=B) *)
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
270 |
Goal "i : I ==> \ |
| 6536 | 271 |
\ (JN i:I. F i) : A ensures B = \ |
|
8251
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8216
diff
changeset
|
272 |
\ ((ALL i:I. F i : (A-B) co (A Un B)) & (EX i:I. F i : A ensures B))"; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
273 |
by (auto_tac (claset(), |
| 7523 | 274 |
simpset() addsimps [ensures_def, JN_constrains, JN_transient])); |
275 |
qed "JN_ensures"; |
|
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
276 |
|
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
277 |
Goalw [ensures_def] |
| 6536 | 278 |
"F Join G : A ensures B = \ |
| 7630 | 279 |
\ (F : (A-B) co (A Un B) & G : (A-B) co (A Un B) & \ |
280 |
\ (F : transient (A-B) | G : transient (A-B)))"; |
|
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
8314
diff
changeset
|
281 |
by (auto_tac (claset(), simpset() addsimps [Join_transient])); |
| 7523 | 282 |
qed "Join_ensures"; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
283 |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
284 |
Goalw [stable_def, constrains_def, Join_def] |
| 6536 | 285 |
"[| F : stable A; G : A co A' |] \ |
286 |
\ ==> F Join G : A co A'"; |
|
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6019
diff
changeset
|
287 |
by (asm_full_simp_tac (simpset() addsimps [ball_Un]) 1); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
288 |
by (Blast_tac 1); |
| 7523 | 289 |
qed "stable_Join_constrains"; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
290 |
|
| 6633 | 291 |
(*Premise for G cannot use Always because F: Stable A is weaker than |
| 5648 | 292 |
G : stable A *) |
| 6570 | 293 |
Goal "[| F : stable A; G : invariant A |] ==> F Join G : Always A"; |
294 |
by (full_simp_tac (simpset() addsimps [Always_def, invariant_def, |
|
|
9403
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
paulson
parents:
8314
diff
changeset
|
295 |
Stable_eq_stable]) 1); |
|
8069
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
paulson
parents:
8055
diff
changeset
|
296 |
by (force_tac(claset() addIs [stable_Int], simpset()) 1); |
| 8216 | 297 |
qed "stable_Join_Always1"; |
298 |
||
299 |
(*As above, but exchanging the roles of F and G*) |
|
300 |
Goal "[| F : invariant A; G : stable A |] ==> F Join G : Always A"; |
|
301 |
by (stac Join_commute 1); |
|
302 |
by (blast_tac (claset() addIs [stable_Join_Always1]) 1); |
|
303 |
qed "stable_Join_Always2"; |
|
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
304 |
|
| 6536 | 305 |
Goal "[| F : stable A; G : A ensures B |] ==> F Join G : A ensures B"; |
| 7523 | 306 |
by (asm_simp_tac (simpset() addsimps [Join_ensures]) 1); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
307 |
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
|
308 |
by (etac constrains_weaken 1); |
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
309 |
by Auto_tac; |
| 7523 | 310 |
qed "stable_Join_ensures1"; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
311 |
|
|
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
312 |
(*As above, but exchanging the roles of F and G*) |
| 6536 | 313 |
Goal "[| F : A ensures B; G : stable A |] ==> F Join G : A ensures B"; |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
314 |
by (stac Join_commute 1); |
| 7523 | 315 |
by (blast_tac (claset() addIs [stable_Join_ensures1]) 1); |
316 |
qed "stable_Join_ensures2"; |
|
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5259
diff
changeset
|
317 |
|
| 5648 | 318 |