| author | nipkow | 
| Wed, 29 Mar 2000 15:09:51 +0200 | |
| changeset 8604 | c99e0024050c | 
| parent 8251 | 9be357df93d4 | 
| child 9403 | aad13b59b8d9 | 
| permissions | -rw-r--r-- | 
| 5597 | 1 | (* Title: HOL/UNITY/Comp.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1998 University of Cambridge | |
| 5 | ||
| 6 | Composition | |
| 7 | ||
| 8 | From Chandy and Sanders, "Reasoning About Program Composition" | |
| 9 | *) | |
| 10 | ||
| 11 | (*** component ***) | |
| 12 | ||
| 6738 
06189132c67b
component_eq_subset: a neat characterization of "component"
 paulson parents: 
6703diff
changeset | 13 | Goalw [component_def] | 
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7386diff
changeset | 14 | "H <= F | H <= G ==> H <= (F Join G)"; | 
| 7361 | 15 | by Auto_tac; | 
| 16 | by (res_inst_tac [("x", "G Join Ga")] exI 1);
 | |
| 17 | by (res_inst_tac [("x", "G Join F")] exI 2);
 | |
| 18 | by (auto_tac (claset(), simpset() addsimps Join_ac)); | |
| 19 | qed "componentI"; | |
| 20 | ||
| 21 | Goalw [component_def] | |
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7386diff
changeset | 22 | "(F <= G) = (Init G <= Init F & Acts F <= Acts G)"; | 
| 6738 
06189132c67b
component_eq_subset: a neat characterization of "component"
 paulson parents: 
6703diff
changeset | 23 | by (force_tac (claset() addSIs [exI, program_equalityI], | 
| 7537 | 24 | simpset()) 1); | 
| 6738 
06189132c67b
component_eq_subset: a neat characterization of "component"
 paulson parents: 
6703diff
changeset | 25 | qed "component_eq_subset"; | 
| 
06189132c67b
component_eq_subset: a neat characterization of "component"
 paulson parents: 
6703diff
changeset | 26 | |
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7386diff
changeset | 27 | Goalw [component_def] "SKIP <= F"; | 
| 6012 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 paulson parents: 
5968diff
changeset | 28 | by (force_tac (claset() addIs [Join_SKIP_left], simpset()) 1); | 
| 5612 
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
 paulson parents: 
5597diff
changeset | 29 | qed "component_SKIP"; | 
| 
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
 paulson parents: 
5597diff
changeset | 30 | |
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7386diff
changeset | 31 | Goalw [component_def] "F <= (F :: 'a program)"; | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6018diff
changeset | 32 | by (blast_tac (claset() addIs [Join_SKIP_right]) 1); | 
| 5597 | 33 | qed "component_refl"; | 
| 34 | ||
| 5612 
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
 paulson parents: 
5597diff
changeset | 35 | AddIffs [component_SKIP, component_refl]; | 
| 5597 | 36 | |
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7386diff
changeset | 37 | Goal "F <= SKIP ==> F = SKIP"; | 
| 6833 | 38 | by (auto_tac (claset() addSIs [program_equalityI], | 
| 39 | simpset() addsimps [component_eq_subset])); | |
| 40 | qed "SKIP_minimal"; | |
| 41 | ||
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7386diff
changeset | 42 | Goalw [component_def] "F <= (F Join G)"; | 
| 5968 | 43 | by (Blast_tac 1); | 
| 44 | qed "component_Join1"; | |
| 45 | ||
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7386diff
changeset | 46 | Goalw [component_def] "G <= (F Join G)"; | 
| 5968 | 47 | by (simp_tac (simpset() addsimps [Join_commute]) 1); | 
| 48 | by (Blast_tac 1); | |
| 49 | qed "component_Join2"; | |
| 50 | ||
| 7915 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 51 | Goal "F<=G ==> F Join G = G"; | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 52 | by (auto_tac (claset(), simpset() addsimps [component_def, Join_left_absorb])); | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 53 | qed "Join_absorb1"; | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 54 | |
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 55 | Goal "G<=F ==> F Join G = F"; | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 56 | by (auto_tac (claset(), simpset() addsimps Join_ac@[component_def])); | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 57 | qed "Join_absorb2"; | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7878diff
changeset | 58 | |
| 7947 | 59 | Goal "((JOIN I F) <= H) = (ALL i: I. F i <= H)"; | 
| 60 | by (simp_tac (simpset() addsimps [component_eq_subset]) 1); | |
| 61 | by (Blast_tac 1); | |
| 62 | qed "JN_component_iff"; | |
| 63 | ||
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7386diff
changeset | 64 | Goalw [component_def] "i : I ==> (F i) <= (JN i:I. (F i))"; | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6018diff
changeset | 65 | by (blast_tac (claset() addIs [JN_absorb]) 1); | 
| 5968 | 66 | qed "component_JN"; | 
| 67 | ||
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7386diff
changeset | 68 | Goalw [component_def] "[| F <= G; G <= H |] ==> F <= (H :: 'a program)"; | 
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6018diff
changeset | 69 | by (blast_tac (claset() addIs [Join_assoc RS sym]) 1); | 
| 5597 | 70 | qed "component_trans"; | 
| 71 | ||
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7386diff
changeset | 72 | Goal "[| F <= G; G <= F |] ==> F = (G :: 'a program)"; | 
| 6738 
06189132c67b
component_eq_subset: a neat characterization of "component"
 paulson parents: 
6703diff
changeset | 73 | by (full_simp_tac (simpset() addsimps [component_eq_subset]) 1); | 
| 
06189132c67b
component_eq_subset: a neat characterization of "component"
 paulson parents: 
6703diff
changeset | 74 | by (blast_tac (claset() addSIs [program_equalityI]) 1); | 
| 6703 | 75 | qed "component_antisym"; | 
| 5597 | 76 | |
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7386diff
changeset | 77 | Goal "((F Join G) <= H) = (F <= H & G <= H)"; | 
| 7537 | 78 | by (simp_tac (simpset() addsimps [component_eq_subset]) 1); | 
| 6833 | 79 | by (Blast_tac 1); | 
| 80 | qed "Join_component_iff"; | |
| 81 | ||
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7386diff
changeset | 82 | Goal "[| F <= G; G : A co B |] ==> F : A co B"; | 
| 7386 | 83 | by (auto_tac (claset(), | 
| 84 | simpset() addsimps [constrains_def, component_eq_subset])); | |
| 85 | qed "component_constrains"; | |
| 6295 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 paulson parents: 
6018diff
changeset | 86 | |
| 7947 | 87 | (*Used in Guar.thy to show that programs are partially ordered*) | 
| 7399 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 paulson parents: 
7386diff
changeset | 88 | bind_thm ("program_less_le", strict_component_def RS meta_eq_to_obj_eq);
 | 
| 7947 | 89 | |
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 90 | |
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 91 | (*** preserves ***) | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 92 | |
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 93 | val prems = | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 94 | Goalw [preserves_def] "(!!z. F : stable {s. v s = z}) ==> F : preserves v";
 | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 95 | by (blast_tac (claset() addIs prems) 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 96 | qed "preservesI"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 97 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 98 | Goalw [preserves_def, stable_def, constrains_def] | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 99 | "[| F : preserves v; act : Acts F; (s,s') : act |] ==> v s = v s'"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 100 | by (Force_tac 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 101 | qed "preserves_imp_eq"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 102 | |
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 103 | Goal "(F Join G : preserves v) = (F : preserves v & G : preserves v)"; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 104 | by (simp_tac (simpset() addsimps [Join_stable, preserves_def]) 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 105 | by (Blast_tac 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 106 | qed "Join_preserves"; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 107 | |
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 108 | Goal "(JOIN I F : preserves v) = (ALL i:I. F i : preserves v)"; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 109 | by (simp_tac (simpset() addsimps [JN_stable, preserves_def]) 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 110 | by (Blast_tac 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 111 | qed "JN_preserves"; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 112 | |
| 8065 
658e0d4e4ed9
first working version to Alloc/System_Client_Progress;
 paulson parents: 
8055diff
changeset | 113 | AddIffs [Join_preserves, JN_preserves]; | 
| 
658e0d4e4ed9
first working version to Alloc/System_Client_Progress;
 paulson parents: 
8055diff
changeset | 114 | |
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 115 | Goalw [funPair_def] "(funPair f g) x = (f x, g x)"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 116 | by (Simp_tac 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 117 | qed "funPair_apply"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 118 | Addsimps [funPair_apply]; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 119 | |
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 120 | Goal "preserves (funPair v w) = preserves v Int preserves w"; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 121 | by (auto_tac (claset(), | 
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 122 | simpset() addsimps [preserves_def, stable_def, constrains_def])); | 
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 123 | by (Blast_tac 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 124 | qed "preserves_funPair"; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 125 | |
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 126 | (* (F : preserves (funPair v w)) = (F : preserves v Int preserves w) *) | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 127 | AddIffs [preserves_funPair RS eqset_imp_iff]; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 128 | |
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 129 | |
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 130 | Goal "(funPair f g) o h = funPair (f o h) (g o h)"; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 131 | by (simp_tac (simpset() addsimps [funPair_def, o_def]) 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 132 | qed "funPair_o_distrib"; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 133 | |
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 134 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 135 | Goal "fst o (funPair f g) = f"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 136 | by (simp_tac (simpset() addsimps [funPair_def, o_def]) 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 137 | qed "fst_o_funPair"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 138 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 139 | Goal "snd o (funPair f g) = g"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 140 | by (simp_tac (simpset() addsimps [funPair_def, o_def]) 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 141 | qed "snd_o_funPair"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 142 | Addsimps [fst_o_funPair, snd_o_funPair]; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 143 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 144 | |
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 145 | Goal "preserves v <= preserves (w o v)"; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 146 | by (force_tac (claset(), | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 147 | simpset() addsimps [preserves_def, | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 148 | stable_def, constrains_def]) 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 149 | qed "subset_preserves_o"; | 
| 7947 | 150 | |
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 151 | Goal "preserves v <= stable {s. P (v s)}";
 | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 152 | by (auto_tac (claset(), | 
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 153 | simpset() addsimps [preserves_def, stable_def, constrains_def])); | 
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 154 | by (rename_tac "s' s" 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 155 | by (subgoal_tac "v s = v s'" 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 156 | by (ALLGOALS Force_tac); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 157 | qed "preserves_subset_stable"; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 158 | |
| 8216 | 159 | Goal "preserves v <= increasing v"; | 
| 160 | by (auto_tac (claset(), | |
| 161 | simpset() addsimps [impOfSubs preserves_subset_stable, | |
| 162 | increasing_def])); | |
| 163 | qed "preserves_subset_increasing"; | |
| 164 | ||
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 165 | Goal "preserves id <= stable A"; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 166 | by (force_tac (claset(), | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 167 | simpset() addsimps [preserves_def, stable_def, constrains_def]) 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 168 | qed "preserves_id_subset_stable"; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 169 | |
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 170 | |
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 171 | (** Some lemmas used only in Client.ML **) | 
| 7947 | 172 | |
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 173 | Goal "[| F : stable {s. P (v s) (w s)};   \
 | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 174 | \ G : preserves v; G : preserves w |] \ | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 175 | \     ==> F Join G : stable {s. P (v s) (w s)}";
 | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 176 | by (asm_simp_tac (simpset() addsimps [Join_stable]) 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 177 | by (subgoal_tac "G: preserves (funPair v w)" 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 178 | by (Asm_simp_tac 2); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 179 | by (dres_inst_tac [("P1", "split ?Q")]  
 | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 180 | (impOfSubs preserves_subset_stable) 1); | 
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 181 | by Auto_tac; | 
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 182 | qed "stable_localTo_stable2"; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 183 | |
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 184 | Goal "[| F : stable {s. v s <= w s};  G : preserves v;       \
 | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 185 | \ F Join G : Increasing w |] \ | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 186 | \     ==> F Join G : Stable {s. v s <= w s}";
 | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 187 | by (auto_tac (claset(), | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 188 | simpset() addsimps [stable_def, Stable_def, Increasing_def, | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 189 | Constrains_def, Join_constrains, all_conj_distrib])); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 190 | by (blast_tac (claset() addIs [constrains_weaken]) 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 191 | (*The G case remains*) | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 192 | by (auto_tac (claset(), | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 193 | simpset() addsimps [preserves_def, stable_def, constrains_def])); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 194 | by (case_tac "act: Acts F" 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 195 | by (Blast_tac 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 196 | (*We have a G-action, so delete assumptions about F-actions*) | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 197 | by (thin_tac "ALL act:Acts F. ?P act" 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 198 | by (thin_tac "ALL z. ALL act:Acts F. ?P z act" 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 199 | by (subgoal_tac "v x = v xa" 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 200 | by (Blast_tac 2); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 201 | by Auto_tac; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 202 | by (etac order_trans 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 203 | by (Blast_tac 1); | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
7947diff
changeset | 204 | qed "Increasing_preserves_Stable"; | 
| 8122 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8065diff
changeset | 205 |