| author | wenzelm | 
| Fri, 03 Nov 2000 21:29:56 +0100 | |
| changeset 10382 | 1fb807260ff1 | 
| parent 10064 | 1a77667b21ef | 
| child 11467 | 1064effe37f6 | 
| 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: 
6019diff
changeset | 11 | |
| 5867 | 12 | (** SKIP **) | 
| 13 | ||
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
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: 
6019diff
changeset | 18 | Goal "Acts SKIP = {Id}";
 | 
| 5648 | 19 | by (simp_tac (simpset() addsimps [SKIP_def]) 1); | 
| 20 | qed "Acts_SKIP"; | |
| 21 | ||
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 22 | Goal "AllowedActs SKIP = UNIV"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 23 | by (auto_tac (claset(), simpset() addsimps [SKIP_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 24 | qed "AllowedActs_SKIP"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 25 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 26 | Addsimps [Init_SKIP, Acts_SKIP, AllowedActs_SKIP]; | 
| 5648 | 27 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 28 | Goal "reachable SKIP = UNIV"; | 
| 6012 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5970diff
changeset | 29 | by (force_tac (claset() addEs [reachable.induct] | 
| 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5970diff
changeset | 30 | addIs reachable.intrs, simpset()) 1); | 
| 5867 | 31 | qed "reachable_SKIP"; | 
| 32 | ||
| 33 | Addsimps [reachable_SKIP]; | |
| 34 | ||
| 6836 | 35 | (** SKIP and safety properties **) | 
| 36 | ||
| 37 | Goalw [constrains_def] "(SKIP : A co B) = (A<=B)"; | |
| 38 | by Auto_tac; | |
| 39 | qed "SKIP_in_constrains_iff"; | |
| 40 | AddIffs [SKIP_in_constrains_iff]; | |
| 41 | ||
| 42 | Goalw [Constrains_def] "(SKIP : A Co B) = (A<=B)"; | |
| 43 | by Auto_tac; | |
| 44 | qed "SKIP_in_Constrains_iff"; | |
| 45 | AddIffs [SKIP_in_Constrains_iff]; | |
| 46 | ||
| 47 | Goalw [stable_def] "SKIP : stable A"; | |
| 48 | by Auto_tac; | |
| 49 | qed "SKIP_in_stable"; | |
| 50 | AddIffs [SKIP_in_stable, SKIP_in_stable RS stable_imp_Stable]; | |
| 51 | ||
| 5867 | 52 | |
| 6012 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5970diff
changeset | 53 | (** Join **) | 
| 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5970diff
changeset | 54 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 55 | Goal "Init (F Join G) = Init F Int Init G"; | 
| 6012 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5970diff
changeset | 56 | by (simp_tac (simpset() addsimps [Join_def]) 1); | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 57 | qed "Init_Join"; | 
| 5252 | 58 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 59 | Goal "Acts (F Join G) = Acts F Un Acts G"; | 
| 5596 | 60 | by (auto_tac (claset(), simpset() addsimps [Join_def])); | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 61 | qed "Acts_Join"; | 
| 5252 | 62 | |
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 63 | Goal "AllowedActs (F Join G) = AllowedActs F Int AllowedActs G"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 64 | by (auto_tac (claset(), simpset() addsimps [Join_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 65 | qed "AllowedActs_Join"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 66 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 67 | Addsimps [Init_Join, Acts_Join, AllowedActs_Join]; | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 68 | |
| 6012 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5970diff
changeset | 69 | |
| 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5970diff
changeset | 70 | (** JN **) | 
| 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5970diff
changeset | 71 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 72 | Goalw [JOIN_def, SKIP_def] "(JN i:{}. F i) = SKIP";
 | 
| 6012 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5970diff
changeset | 73 | by Auto_tac; | 
| 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5970diff
changeset | 74 | qed "JN_empty"; | 
| 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5970diff
changeset | 75 | Addsimps [JN_empty]; | 
| 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5970diff
changeset | 76 | |
| 6633 | 77 | 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: 
6019diff
changeset | 78 | by (rtac program_equalityI 1); | 
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 79 | by (auto_tac (claset(), simpset() addsimps [JOIN_def, Join_def])); | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 80 | qed "JN_insert"; | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 81 | Addsimps[JN_empty, JN_insert]; | 
| 6012 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5970diff
changeset | 82 | |
| 5611 | 83 | Goal "Init (JN i:I. F i) = (INT i:I. Init (F i))"; | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 84 | by (simp_tac (simpset() addsimps [JOIN_def]) 1); | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 85 | qed "Init_JN"; | 
| 5252 | 86 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 87 | Goal "Acts (JN i:I. F i) = insert Id (UN i:I. Acts (F i))"; | 
| 5596 | 88 | by (auto_tac (claset(), simpset() addsimps [JOIN_def])); | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 89 | qed "Acts_JN"; | 
| 5252 | 90 | |
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 91 | Goal "AllowedActs (JN i:I. F i) = (INT i:I. AllowedActs (F i))"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 92 | by (auto_tac (claset(), simpset() addsimps [JOIN_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 93 | qed "AllowedActs_JN"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 94 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 95 | Addsimps [Init_JN, Acts_JN, AllowedActs_JN]; | 
| 5867 | 96 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 97 | val prems = Goalw [JOIN_def] | 
| 6633 | 98 | "[| I=J; !!i. i:J ==> F i = G i |] ==> \ | 
| 99 | \ (JN i:I. F i) = (JN i:J. G i)"; | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 100 | by (asm_simp_tac (simpset() addsimps prems) 1); | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 101 | qed "JN_cong"; | 
| 6012 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5970diff
changeset | 102 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 103 | Addcongs [JN_cong]; | 
| 5584 | 104 | |
| 105 | ||
| 5611 | 106 | (** Algebraic laws **) | 
| 5259 | 107 | |
| 5611 | 108 | Goal "F Join G = G Join F"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 109 | by (simp_tac (simpset() addsimps [Join_def, Un_commute, Int_commute]) 1); | 
| 5259 | 110 | qed "Join_commute"; | 
| 111 | ||
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 112 | |
| 7360 | 113 | Goal "A Join (B Join C) = B Join (A Join C)"; | 
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 114 | by (simp_tac (simpset() addsimps Un_ac@Int_ac@[Join_def, insert_absorb]) 1); | 
| 7360 | 115 | qed "Join_left_commute"; | 
| 116 | ||
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 117 | |
| 5611 | 118 | Goal "(F Join G) Join H = F Join (G Join H)"; | 
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 119 | by (simp_tac (simpset() addsimps Un_ac@[Join_def, Int_assoc, insert_absorb]) 1); | 
| 5259 | 120 | qed "Join_assoc"; | 
| 5596 | 121 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 122 | Goalw [Join_def, SKIP_def] "SKIP Join F = F"; | 
| 5611 | 123 | by (rtac program_equalityI 1); | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 124 | by (ALLGOALS (simp_tac (simpset() addsimps [insert_absorb]))); | 
| 5611 | 125 | qed "Join_SKIP_left"; | 
| 5584 | 126 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 127 | Goalw [Join_def, SKIP_def] "F Join SKIP = F"; | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 128 | by (rtac program_equalityI 1); | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 129 | by (ALLGOALS (simp_tac (simpset() addsimps [insert_absorb]))); | 
| 5611 | 130 | qed "Join_SKIP_right"; | 
| 131 | ||
| 132 | Addsimps [Join_SKIP_left, Join_SKIP_right]; | |
| 133 | ||
| 134 | Goalw [Join_def] "F Join F = F"; | |
| 135 | by (rtac program_equalityI 1); | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 136 | by Auto_tac; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 137 | qed "Join_absorb"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 138 | |
| 5611 | 139 | Addsimps [Join_absorb]; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 140 | |
| 7915 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 141 | Goalw [Join_def] "F Join (F Join G) = F Join G"; | 
| 7360 | 142 | by (rtac program_equalityI 1); | 
| 143 | by Auto_tac; | |
| 144 | qed "Join_left_absorb"; | |
| 145 | ||
| 146 | (*Join is an AC-operator*) | |
| 147 | val Join_ac = [Join_assoc, Join_left_absorb, Join_commute, Join_left_commute]; | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 148 | |
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 149 | |
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 150 | (*** JN laws ***) | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 151 | |
| 6633 | 152 | (*Also follows by JN_insert and insert_absorb, but the proof is longer*) | 
| 153 | Goal "k:I ==> F k Join (JN i:I. F i) = (JN i:I. F i)"; | |
| 7537 | 154 | by (auto_tac (claset() addSIs [program_equalityI], simpset())); | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 155 | qed "JN_absorb"; | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 156 | |
| 6633 | 157 | Goal "(JN i: I Un J. F i) = ((JN i: I. F i) Join (JN i:J. F i))"; | 
| 7537 | 158 | by (auto_tac (claset() addSIs [program_equalityI], simpset())); | 
| 6633 | 159 | qed "JN_Un"; | 
| 5970 | 160 | |
| 6836 | 161 | Goal "(JN i:I. c) = (if I={} then SKIP else c)";
 | 
| 8314 | 162 | by (rtac program_equalityI 1); | 
| 163 | by Auto_tac; | |
| 5970 | 164 | qed "JN_constant"; | 
| 165 | ||
| 6633 | 166 | Goal "(JN i:I. F i Join G i) = (JN i:I. F i) Join (JN i:I. G i)"; | 
| 7537 | 167 | by (auto_tac (claset() addSIs [program_equalityI], simpset())); | 
| 5970 | 168 | qed "JN_Join_distrib"; | 
| 5584 | 169 | |
| 6633 | 170 | Goal "i : I ==> (JN i:I. F i Join G) = ((JN i:I. F i) Join G)"; | 
| 171 | by (asm_simp_tac (simpset() addsimps [JN_Join_distrib, JN_constant]) 1); | |
| 6836 | 172 | by Auto_tac; | 
| 6633 | 173 | qed "JN_Join_miniscope"; | 
| 174 | ||
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
8042diff
changeset | 175 | (*Used to prove guarantees_JN_I*) | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
8042diff
changeset | 176 | 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: 
8042diff
changeset | 177 | by (rtac program_equalityI 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
8042diff
changeset | 178 | by Auto_tac; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
8042diff
changeset | 179 | qed "JN_Join_diff"; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
8042diff
changeset | 180 | |
| 5584 | 181 | |
| 6536 | 182 | (*** Safety: co, stable, FP ***) | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 183 | |
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 184 | (*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: 
8216diff
changeset | 185 | alternative precondition is A<=B, but most proofs using this rule require | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 186 | I to be nonempty for other reasons anyway.*) | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 187 | Goalw [constrains_def, JOIN_def] | 
| 6633 | 188 | "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: 
6019diff
changeset | 189 | by (Simp_tac 1); | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 190 | by (Blast_tac 1); | 
| 7523 | 191 | qed "JN_constrains"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 192 | |
| 6633 | 193 | 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: 
6019diff
changeset | 194 | by (auto_tac | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 195 | (claset(), | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 196 | simpset() addsimps [constrains_def, Join_def])); | 
| 7523 | 197 | qed "Join_constrains"; | 
| 5620 | 198 | |
| 8122 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8069diff
changeset | 199 | 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: 
8069diff
changeset | 200 | 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: 
8069diff
changeset | 201 | qed "Join_unless"; | 
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8069diff
changeset | 202 | |
| 9403 
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
 paulson parents: 
8314diff
changeset | 203 | Addsimps [Join_constrains, Join_unless]; | 
| 
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
 paulson parents: 
8314diff
changeset | 204 | |
| 6633 | 205 | (*Analogous weak versions FAIL; see Misra [1994] 5.4.1, Substitution Axiom. | 
| 206 | reachable (F Join G) could be much bigger than reachable F, reachable G | |
| 207 | *) | |
| 6012 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5970diff
changeset | 208 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 209 | |
| 6536 | 210 | Goal "[| F : A co A'; G : B co B' |] \ | 
| 211 | \ ==> F Join G : (A Int B) co (A' Un B')"; | |
| 9403 
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
 paulson parents: 
8314diff
changeset | 212 | by (Simp_tac 1); | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 213 | by (blast_tac (claset() addIs [constrains_weaken]) 1); | 
| 7523 | 214 | qed "Join_constrains_weaken"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 215 | |
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 216 | (*If I={}, it degenerates to SKIP : UNIV co {}, which is false.*)
 | 
| 6633 | 217 | Goal "[| ALL i:I. F i : A i co A' i; i: I |] \ | 
| 218 | \ ==> (JN i:I. F i) : (INT i:I. A i) co (UN i:I. A' i)"; | |
| 7523 | 219 | by (asm_simp_tac (simpset() addsimps [JN_constrains]) 1); | 
| 6633 | 220 | by (blast_tac (claset() addIs [constrains_weaken]) 1); | 
| 7523 | 221 | qed "JN_constrains_weaken"; | 
| 6633 | 222 | |
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
8042diff
changeset | 223 | 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: 
8042diff
changeset | 224 | by (asm_simp_tac | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
8042diff
changeset | 225 | (simpset() addsimps [stable_def, constrains_def, JOIN_def]) 1); | 
| 7523 | 226 | qed "JN_stable"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 227 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 228 | Goal "[| ALL i:I. F i : invariant A; i : I |] \ | 
| 5970 | 229 | \ ==> (JN i:I. F i) : invariant A"; | 
| 7523 | 230 | by (asm_full_simp_tac (simpset() addsimps [invariant_def, JN_stable]) 1); | 
| 5970 | 231 | by (Blast_tac 1); | 
| 232 | bind_thm ("invariant_JN_I", ballI RS result());
 | |
| 233 | ||
| 7594 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7537diff
changeset | 234 | Goal "(F Join G : stable A) = \ | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 235 | \ (F : stable A & G : stable A)"; | 
| 9403 
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
 paulson parents: 
8314diff
changeset | 236 | by (simp_tac (simpset() addsimps [stable_def]) 1); | 
| 7523 | 237 | qed "Join_stable"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 238 | |
| 7594 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7537diff
changeset | 239 | Goal "(F Join G : increasing f) = \ | 
| 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7537diff
changeset | 240 | \ (F : increasing f & G : increasing f)"; | 
| 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7537diff
changeset | 241 | by (simp_tac (simpset() addsimps [increasing_def, Join_stable]) 1); | 
| 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7537diff
changeset | 242 | by (Blast_tac 1); | 
| 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7537diff
changeset | 243 | qed "Join_increasing"; | 
| 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7537diff
changeset | 244 | |
| 9403 
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
 paulson parents: 
8314diff
changeset | 245 | Addsimps [Join_stable, Join_increasing]; | 
| 
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
 paulson parents: 
8314diff
changeset | 246 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 247 | Goal "[| F : invariant A; G : invariant A |] \ | 
| 5970 | 248 | \ ==> F Join G : invariant A"; | 
| 9403 
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
 paulson parents: 
8314diff
changeset | 249 | by (full_simp_tac (simpset() addsimps [invariant_def]) 1); | 
| 5970 | 250 | by (Blast_tac 1); | 
| 251 | qed "invariant_JoinI"; | |
| 252 | ||
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
8042diff
changeset | 253 | Goal "FP (JN i:I. F i) = (INT i:I. FP (F i))"; | 
| 7523 | 254 | 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: 
5259diff
changeset | 255 | qed "FP_JN"; | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 256 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 257 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 258 | (*** Progress: transient, ensures ***) | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 259 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 260 | Goal "i : I ==> \ | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 261 | \ (JN i:I. F i) : transient A = (EX i:I. F i : transient A)"; | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 262 | by (auto_tac (claset(), | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 263 | simpset() addsimps [transient_def, JOIN_def])); | 
| 7523 | 264 | qed "JN_transient"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 265 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 266 | Goal "F Join G : transient A = \ | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 267 | \ (F : transient A | G : transient A)"; | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 268 | by (auto_tac (claset(), | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 269 | simpset() addsimps [bex_Un, transient_def, | 
| 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 270 | Join_def])); | 
| 7523 | 271 | qed "Join_transient"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 272 | |
| 9403 
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
 paulson parents: 
8314diff
changeset | 273 | Addsimps [Join_transient]; | 
| 
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
 paulson parents: 
8314diff
changeset | 274 | |
| 8216 | 275 | Goal "F : transient A ==> F Join G : transient A"; | 
| 276 | by (asm_simp_tac (simpset() addsimps [Join_transient]) 1); | |
| 277 | qed "Join_transient_I1"; | |
| 278 | ||
| 279 | Goal "G : transient A ==> F Join G : transient A"; | |
| 280 | by (asm_simp_tac (simpset() addsimps [Join_transient]) 1); | |
| 281 | qed "Join_transient_I2"; | |
| 282 | ||
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 283 | (*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: 
6019diff
changeset | 284 | Goal "i : I ==> \ | 
| 6536 | 285 | \ (JN i:I. F i) : A ensures B = \ | 
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 286 | \ ((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: 
5259diff
changeset | 287 | by (auto_tac (claset(), | 
| 7523 | 288 | simpset() addsimps [ensures_def, JN_constrains, JN_transient])); | 
| 289 | qed "JN_ensures"; | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 290 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 291 | Goalw [ensures_def] | 
| 6536 | 292 | "F Join G : A ensures B = \ | 
| 7630 | 293 | \ (F : (A-B) co (A Un B) & G : (A-B) co (A Un B) & \ | 
| 294 | \ (F : transient (A-B) | G : transient (A-B)))"; | |
| 9403 
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
 paulson parents: 
8314diff
changeset | 295 | by (auto_tac (claset(), simpset() addsimps [Join_transient])); | 
| 7523 | 296 | qed "Join_ensures"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 297 | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 298 | Goalw [stable_def, constrains_def, Join_def] | 
| 6536 | 299 | "[| F : stable A; G : A co A' |] \ | 
| 300 | \ ==> F Join G : A co A'"; | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6019diff
changeset | 301 | by (asm_full_simp_tac (simpset() addsimps [ball_Un]) 1); | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 302 | by (Blast_tac 1); | 
| 7523 | 303 | qed "stable_Join_constrains"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 304 | |
| 6633 | 305 | (*Premise for G cannot use Always because F: Stable A is weaker than | 
| 5648 | 306 | G : stable A *) | 
| 6570 | 307 | Goal "[| F : stable A; G : invariant A |] ==> F Join G : Always A"; | 
| 308 | by (full_simp_tac (simpset() addsimps [Always_def, invariant_def, | |
| 9403 
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
 paulson parents: 
8314diff
changeset | 309 | Stable_eq_stable]) 1); | 
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 310 | by (force_tac(claset() addIs [stable_Int], simpset()) 1); | 
| 8216 | 311 | qed "stable_Join_Always1"; | 
| 312 | ||
| 313 | (*As above, but exchanging the roles of F and G*) | |
| 314 | Goal "[| F : invariant A; G : stable A |] ==> F Join G : Always A"; | |
| 315 | by (stac Join_commute 1); | |
| 316 | by (blast_tac (claset() addIs [stable_Join_Always1]) 1); | |
| 317 | qed "stable_Join_Always2"; | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 318 | |
| 6536 | 319 | Goal "[| F : stable A; G : A ensures B |] ==> F Join G : A ensures B"; | 
| 7523 | 320 | by (asm_simp_tac (simpset() addsimps [Join_ensures]) 1); | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 321 | 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: 
5259diff
changeset | 322 | by (etac constrains_weaken 1); | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 323 | by Auto_tac; | 
| 7523 | 324 | qed "stable_Join_ensures1"; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 325 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 326 | (*As above, but exchanging the roles of F and G*) | 
| 6536 | 327 | 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: 
5259diff
changeset | 328 | by (stac Join_commute 1); | 
| 7523 | 329 | by (blast_tac (claset() addIs [stable_Join_ensures1]) 1); | 
| 330 | qed "stable_Join_ensures2"; | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5259diff
changeset | 331 | |
| 5648 | 332 | |
| 10064 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 333 | (*** the ok and OK relations ***) | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 334 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 335 | Goal "SKIP ok F"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 336 | by (auto_tac (claset(), simpset() addsimps [ok_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 337 | qed "ok_SKIP1"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 338 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 339 | Goal "F ok SKIP"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 340 | by (auto_tac (claset(), simpset() addsimps [ok_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 341 | qed "ok_SKIP2"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 342 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 343 | AddIffs [ok_SKIP1, ok_SKIP2]; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 344 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 345 | Goal "(F ok G & (F Join G) ok H) = (G ok H & F ok (G Join H))"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 346 | by (auto_tac (claset(), simpset() addsimps [ok_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 347 | qed "ok_Join_commute"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 348 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 349 | Goal "(F ok G) = (G ok F)"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 350 | by (auto_tac (claset(), simpset() addsimps [ok_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 351 | qed "ok_commute"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 352 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 353 | bind_thm ("ok_sym", ok_commute RS iffD1);
 | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 354 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 355 | Goal "OK {(0,F),(1,G),(2,H)} snd = (F ok G & (F Join G) ok H)";
 | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 356 | by (asm_full_simp_tac | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 357 | (simpset() addsimps [Ball_def, conj_disj_distribR, ok_def, Join_def, | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 358 | OK_def, insert_absorb, all_conj_distrib, eq_commute]) 1); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 359 | by (Blast_tac 1); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 360 | qed "ok_iff_OK"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 361 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 362 | Goal "F ok (G Join H) = (F ok G & F ok H)"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 363 | by (auto_tac (claset(), simpset() addsimps [ok_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 364 | qed "ok_Join_iff1"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 365 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 366 | Goal "(G Join H) ok F = (G ok F & H ok F)"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 367 | by (auto_tac (claset(), simpset() addsimps [ok_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 368 | qed "ok_Join_iff2"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 369 | AddIffs [ok_Join_iff1, ok_Join_iff2]; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 370 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 371 | (*useful? Not with the previous two around*) | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 372 | Goal "[| F ok G; (F Join G) ok H |] ==> F ok (G Join H)"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 373 | by (auto_tac (claset(), simpset() addsimps [ok_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 374 | qed "ok_Join_commute_I"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 375 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 376 | Goal "F ok (JOIN I G) = (ALL i:I. F ok G i)"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 377 | by (auto_tac (claset(), simpset() addsimps [ok_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 378 | qed "ok_JN_iff1"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 379 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 380 | Goal "(JOIN I G) ok F = (ALL i:I. G i ok F)"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 381 | by (auto_tac (claset(), simpset() addsimps [ok_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 382 | qed "ok_JN_iff2"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 383 | AddIffs [ok_JN_iff1, ok_JN_iff2]; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 384 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 385 | Goal "OK I F = (ALL i: I. ALL j: I-{i}. (F i) ok (F j))"; 
 | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 386 | by (auto_tac (claset(), simpset() addsimps [ok_def, OK_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 387 | qed "OK_iff_ok"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 388 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 389 | Goal "[| OK I F; i: I; j: I; i ~= j|] ==> (F i) ok (F j)"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 390 | by (auto_tac (claset(), simpset() addsimps [OK_iff_ok])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 391 | qed "OK_imp_ok"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 392 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 393 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 394 | (*** Allowed ***) | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 395 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 396 | Goal "Allowed SKIP = UNIV"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 397 | by (auto_tac (claset(), simpset() addsimps [Allowed_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 398 | qed "Allowed_SKIP"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 399 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 400 | Goal "Allowed (F Join G) = Allowed F Int Allowed G"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 401 | by (auto_tac (claset(), simpset() addsimps [Allowed_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 402 | qed "Allowed_Join"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 403 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 404 | Goal "Allowed (JOIN I F) = (INT i:I. Allowed (F i))"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 405 | by (auto_tac (claset(), simpset() addsimps [Allowed_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 406 | qed "Allowed_JN"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 407 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 408 | Addsimps [Allowed_SKIP, Allowed_Join, Allowed_JN]; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 409 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 410 | Goal "F ok G = (F : Allowed G & G : Allowed F)"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 411 | by (simp_tac (simpset() addsimps [ok_def, Allowed_def]) 1); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 412 | qed "ok_iff_Allowed"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 413 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 414 | Goal "OK I F = (ALL i: I. ALL j: I-{i}. F i : Allowed(F j))"; 
 | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 415 | by (auto_tac (claset(), simpset() addsimps [OK_iff_ok, ok_iff_Allowed])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 416 | qed "OK_iff_Allowed"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 417 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 418 | (*** safety_prop, for reasoning about given instances of "ok" ***) | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 419 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 420 | Goal "safety_prop X ==> (Acts G <= insert Id (UNION X Acts)) = (G : X)"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 421 | by (auto_tac (claset(), simpset() addsimps [safety_prop_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 422 | qed "safety_prop_Acts_iff"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 423 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 424 | Goal "safety_prop X ==> (UNION X Acts <= AllowedActs F) = (X <= Allowed F)"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 425 | by (auto_tac (claset(), | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 426 | simpset() addsimps [Allowed_def, safety_prop_Acts_iff RS sym])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 427 | qed "safety_prop_AllowedActs_iff_Allowed"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 428 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 429 | Goal "safety_prop X ==> Allowed (mk_program (init, acts, UNION X Acts)) = X"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 430 | by (asm_simp_tac (simpset() addsimps [Allowed_def, safety_prop_Acts_iff]) 1); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 431 | qed "Allowed_eq"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 432 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 433 | Goal "[| F == mk_program (init, acts, UNION X Acts) ; safety_prop X |] \ | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 434 | \ ==> Allowed F = X"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 435 | by (asm_simp_tac (simpset() addsimps [Allowed_eq]) 1); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 436 | qed "def_prg_Allowed"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 437 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 438 | (*For safety_prop to hold, the property must be satisfiable!*) | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 439 | Goal "safety_prop (A co B) = (A <= B)"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 440 | by (simp_tac (simpset() addsimps [safety_prop_def, constrains_def]) 1); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 441 | by (Blast_tac 1); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 442 | qed "safety_prop_constrains"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 443 | AddIffs [safety_prop_constrains]; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 444 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 445 | Goal "safety_prop (stable A)"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 446 | by (simp_tac (simpset() addsimps [stable_def]) 1); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 447 | qed "safety_prop_stable"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 448 | AddIffs [safety_prop_stable]; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 449 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 450 | Goal "[| safety_prop X; safety_prop Y |] ==> safety_prop (X Int Y)"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 451 | by (full_simp_tac (simpset() addsimps [safety_prop_def]) 1); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 452 | by (Blast_tac 1); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 453 | qed "safety_prop_Int"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 454 | Addsimps [safety_prop_Int]; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 455 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 456 | Goal "(ALL i. safety_prop (X i)) ==> safety_prop (INT i. X i)"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 457 | by (auto_tac (claset(), simpset() addsimps [safety_prop_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 458 | by (Blast_tac 1); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 459 | bind_thm ("safety_prop_INTER1", allI RS result());
 | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 460 | Addsimps [safety_prop_INTER1]; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 461 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 462 | Goal "(ALL i:I. safety_prop (X i)) ==> safety_prop (INT i:I. X i)"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 463 | by (auto_tac (claset(), simpset() addsimps [safety_prop_def])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 464 | by (Blast_tac 1); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 465 | bind_thm ("safety_prop_INTER", ballI RS result());
 | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 466 | Addsimps [safety_prop_INTER]; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 467 | |
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 468 | Goal "[| F == mk_program(init,acts,UNION X Acts); safety_prop X |] \ | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 469 | \ ==> F ok G = (G : X & acts <= AllowedActs G)"; | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 470 | by (auto_tac (claset(), simpset() addsimps [ok_def, safety_prop_Acts_iff])); | 
| 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 paulson parents: 
9687diff
changeset | 471 | qed "def_UNION_ok_iff"; |