| author | wenzelm | 
| Mon, 18 Sep 2000 14:49:51 +0200 | |
| changeset 10017 | e146bbfc38c1 | 
| parent 9969 | 4753185f1dd2 | 
| child 10064 | 1a77667b21ef | 
| permissions | -rw-r--r-- | 
| 6297 | 1 | (* Title: HOL/UNITY/Extend.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1999 University of Cambridge | |
| 5 | ||
| 6 | Extending of state sets | |
| 7 | function f (forget) maps the extended state to the original state | |
| 8 | function g (forgotten) maps the extended state to the "extending part" | |
| 7482 | 9 | *) | 
| 7362 
f08fade5ea0d
new laws; changed "guar" back to "guarantees" (sorry)
 paulson parents: 
7341diff
changeset | 10 | |
| 7482 | 11 | (** These we prove OUTSIDE the locale. **) | 
| 12 | ||
| 8948 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 13 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 14 | (*** Restrict [MOVE to Relation.thy?] ***) | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 15 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 16 | Goalw [Restrict_def] "((x,y): Restrict A r) = ((x,y): r & x: A)"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 17 | by (Blast_tac 1); | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 18 | qed "Restrict_iff"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 19 | AddIffs [Restrict_iff]; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 20 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 21 | Goal "Restrict UNIV = id"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 22 | by (rtac ext 1); | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 23 | by (auto_tac (claset(), simpset() addsimps [Restrict_def])); | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 24 | qed "Restrict_UNIV"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 25 | Addsimps [Restrict_UNIV]; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 26 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 27 | Goal "Restrict {} r = {}";
 | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 28 | by (auto_tac (claset(), simpset() addsimps [Restrict_def])); | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 29 | qed "Restrict_empty"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 30 | Addsimps [Restrict_empty]; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 31 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 32 | Goalw [Restrict_def] "Restrict A (Restrict B r) = Restrict (A Int B) r"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 33 | by (Blast_tac 1); | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 34 | qed "Restrict_Int"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 35 | Addsimps [Restrict_Int]; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 36 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 37 | Goalw [Restrict_def] "Domain r <= A ==> Restrict A r = r"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 38 | by Auto_tac; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 39 | qed "Restrict_triv"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 40 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 41 | Goalw [Restrict_def] "Restrict A r <= r"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 42 | by Auto_tac; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 43 | qed "Restrict_subset"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 44 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 45 | Goalw [Restrict_def] | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 46 | "[| A <= B; Restrict B r = Restrict B s |] \ | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 47 | \ ==> Restrict A r = Restrict A s"; | 
| 9190 | 48 | by (Blast_tac 1); | 
| 8948 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 49 | qed "Restrict_eq_mono"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 50 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 51 | Goalw [Restrict_def, image_def] | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 52 | "[| s : RR; Restrict A r = Restrict A s |] \ | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 53 | \ ==> Restrict A r : Restrict A `` RR"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 54 | by Auto_tac; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 55 | qed "Restrict_imageI"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 56 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 57 | Goal "Domain (Restrict A r) = A Int Domain r"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 58 | by (Blast_tac 1); | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 59 | qed "Domain_Restrict"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 60 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 61 | Goal "(Restrict A r) ^^ B = r ^^ (A Int B)"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 62 | by (Blast_tac 1); | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 63 | qed "Image_Restrict"; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 64 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 65 | Addsimps [Domain_Restrict, Image_Restrict]; | 
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 66 | |
| 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 paulson parents: 
8251diff
changeset | 67 | |
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 68 | Goal "f Id = Id ==> insert Id (f``Acts F) = f `` Acts F"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 69 | by (blast_tac (claset() addIs [sym RS image_eqI]) 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 70 | qed "insert_Id_image_Acts"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 71 | |
| 7482 | 72 | (*Possibly easier than reasoning about "inv h"*) | 
| 73 | val [surj_h,prem] = | |
| 74 | Goalw [good_map_def] | |
| 75 | "[| surj h; !! x x' y y'. h(x,y) = h(x',y') ==> x=x' |] ==> good_map h"; | |
| 76 | by (safe_tac (claset() addSIs [surj_h])); | |
| 77 | by (rtac prem 1); | |
| 78 | by (stac (surjective_pairing RS sym) 1); | |
| 79 | by (stac (surj_h RS surj_f_inv_f) 1); | |
| 80 | by (rtac refl 1); | |
| 81 | qed "good_mapI"; | |
| 82 | ||
| 83 | Goalw [good_map_def] "good_map h ==> surj h"; | |
| 84 | by Auto_tac; | |
| 85 | qed "good_map_is_surj"; | |
| 86 | ||
| 87 | (*A convenient way of finding a closed form for inv h*) | |
| 88 | val [surj,prem] = Goalw [inv_def] | |
| 89 | "[| surj h; !! x y. g (h(x,y)) = x |] ==> fst (inv h z) = g z"; | |
| 90 | by (res_inst_tac [("y1", "z")] (surj RS surjD RS exE) 1);
 | |
| 9969 | 91 | by (rtac someI2 1); | 
| 7482 | 92 | by (dres_inst_tac [("f", "g")] arg_cong 2);
 | 
| 93 | by (auto_tac (claset(), simpset() addsimps [prem])); | |
| 94 | qed "fst_inv_equalityI"; | |
| 95 | ||
| 6297 | 96 | |
| 97 | Open_locale "Extend"; | |
| 98 | ||
| 99 | val slice_def = thm "slice_def"; | |
| 100 | ||
| 101 | (*** Trivial properties of f, g, h ***) | |
| 102 | ||
| 7482 | 103 | val good_h = rewrite_rule [good_map_def] (thm "good_h"); | 
| 104 | val surj_h = good_h RS conjunct1; | |
| 6297 | 105 | |
| 106 | val f_def = thm "f_def"; | |
| 107 | val g_def = thm "g_def"; | |
| 108 | ||
| 109 | Goal "f(h(x,y)) = x"; | |
| 7482 | 110 | by (simp_tac (simpset() addsimps [f_def, good_h RS conjunct2]) 1); | 
| 6297 | 111 | qed "f_h_eq"; | 
| 112 | Addsimps [f_h_eq]; | |
| 113 | ||
| 7482 | 114 | Goal "h(x,y) = h(x',y') ==> x=x'"; | 
| 115 | by (dres_inst_tac [("f", "fst o inv h")] arg_cong 1);
 | |
| 116 | (*FIXME: If locales worked properly we could put just "f" above*) | |
| 117 | by (full_simp_tac (simpset() addsimps [f_def, good_h RS conjunct2]) 1); | |
| 118 | qed "h_inject1"; | |
| 7878 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 119 | AddDs [h_inject1]; | 
| 6297 | 120 | |
| 121 | Goal "h(f z, g z) = z"; | |
| 9345 | 122 | by (simp_tac (simpset() addsimps [f_def, g_def, surj_h RS surj_f_inv_f]) 1); | 
| 6297 | 123 | qed "h_f_g_eq"; | 
| 124 | ||
| 7878 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 125 | |
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 126 | (** A sequence of proof steps borrowed from Provers/split_paired_all.ML **) | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 127 | |
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 128 | val cT = TFree ("'c", ["term"]);
 | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 129 | |
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 130 | (* "PROP P x == PROP P (h (f x, g x))" *) | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 131 | val lemma1 = Thm.combination | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 132 |   (Thm.reflexive (cterm_of (sign_of thy) (Free ("P", cT --> propT))))
 | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 133 | (Drule.unvarify (h_f_g_eq RS sym RS eq_reflection)); | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 134 | |
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 135 | val prems = Goalw [lemma1] "(!!u y. PROP P (h (u, y))) ==> PROP P x"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 136 | by (resolve_tac prems 1); | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 137 | val lemma2 = result(); | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 138 | |
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 139 | val prems = Goal "(!!u y. PROP P (h (u, y))) ==> (!!z. PROP P z)"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 140 | by (rtac lemma2 1); | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 141 | by (resolve_tac prems 1); | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 142 | val lemma3 = result(); | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 143 | |
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 144 | val prems = Goal "(!!z. PROP P z) ==> (!!u y. PROP P (h (u, y)))"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 145 | (*not needed for proof, but prevents generalization over h, 'c, etc.*) | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 146 | by (cut_facts_tac [surj_h] 1); | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 147 | by (resolve_tac prems 1); | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 148 | val lemma4 = result(); | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 149 | |
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 150 | val split_extended_all = Thm.equal_intr lemma4 lemma3; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 151 | |
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 152 | |
| 6297 | 153 | (*** extend_set: basic properties ***) | 
| 154 | ||
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 155 | Goal "(x : project_set h C) = (EX y. h(x,y) : C)"; | 
| 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 156 | by (simp_tac (simpset() addsimps [project_set_def]) 1); | 
| 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 157 | qed "project_set_iff"; | 
| 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 158 | |
| 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 159 | AddIffs [project_set_iff]; | 
| 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 160 | |
| 7594 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7546diff
changeset | 161 | Goalw [extend_set_def] "A<=B ==> extend_set h A <= extend_set h B"; | 
| 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7546diff
changeset | 162 | by (Blast_tac 1); | 
| 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7546diff
changeset | 163 | qed "extend_set_mono"; | 
| 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7546diff
changeset | 164 | |
| 7826 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 165 | Goalw [extend_set_def] "z : extend_set h A = (f z : A)"; | 
| 7341 | 166 | by (force_tac (claset() addIs [h_f_g_eq RS sym], simpset()) 1); | 
| 6297 | 167 | qed "mem_extend_set_iff"; | 
| 168 | ||
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 169 | AddIffs [mem_extend_set_iff]; | 
| 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 170 | |
| 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 171 | Goalw [extend_set_def] "(extend_set h A <= extend_set h B) = (A <= B)"; | 
| 7826 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 172 | by (Force_tac 1); | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 173 | qed "extend_set_strict_mono"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 174 | AddIffs [extend_set_strict_mono]; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 175 | |
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 176 | Goalw [extend_set_def] "extend_set h {} = {}";
 | 
| 7378 | 177 | by Auto_tac; | 
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 178 | qed "extend_set_empty"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 179 | Addsimps [extend_set_empty]; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 180 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 181 | Goal "extend_set h {s. P s} = {s. P (f s)}";
 | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 182 | by Auto_tac; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 183 | qed "extend_set_eq_Collect"; | 
| 7378 | 184 | |
| 7594 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7546diff
changeset | 185 | Goal "extend_set h {x} = {s. f s = x}";
 | 
| 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7546diff
changeset | 186 | by Auto_tac; | 
| 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7546diff
changeset | 187 | qed "extend_set_sing"; | 
| 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 paulson parents: 
7546diff
changeset | 188 | |
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 189 | Goalw [extend_set_def] "project_set h (extend_set h C) = C"; | 
| 7341 | 190 | by Auto_tac; | 
| 191 | qed "extend_set_inverse"; | |
| 192 | Addsimps [extend_set_inverse]; | |
| 193 | ||
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 194 | Goalw [extend_set_def] "C <= extend_set h (project_set h C)"; | 
| 7947 | 195 | by (auto_tac (claset(), | 
| 196 | simpset() addsimps [split_extended_all])); | |
| 197 | by (Blast_tac 1); | |
| 198 | qed "extend_set_project_set"; | |
| 199 | ||
| 6297 | 200 | Goal "inj (extend_set h)"; | 
| 7341 | 201 | by (rtac inj_on_inverseI 1); | 
| 202 | by (rtac extend_set_inverse 1); | |
| 6297 | 203 | qed "inj_extend_set"; | 
| 204 | ||
| 7689 | 205 | Goalw [extend_set_def] "extend_set h UNIV = UNIV"; | 
| 7878 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 206 | by (auto_tac (claset(), | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 207 | simpset() addsimps [split_extended_all])); | 
| 7689 | 208 | qed "extend_set_UNIV_eq"; | 
| 209 | Addsimps [standard extend_set_UNIV_eq]; | |
| 210 | ||
| 7482 | 211 | (*** project_set: basic properties ***) | 
| 212 | ||
| 213 | (*project_set is simply image!*) | |
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 214 | Goal "project_set h C = f `` C"; | 
| 7878 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 215 | by (auto_tac (claset() addIs [f_h_eq RS sym], | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 216 | simpset() addsimps [split_extended_all])); | 
| 7482 | 217 | qed "project_set_eq"; | 
| 218 | ||
| 219 | (*Converse appears to fail*) | |
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 220 | Goal "!!z. z : C ==> f z : project_set h C"; | 
| 7878 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 221 | by (auto_tac (claset(), | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 222 | simpset() addsimps [split_extended_all])); | 
| 7482 | 223 | qed "project_set_I"; | 
| 224 | ||
| 225 | ||
| 226 | (*** More laws ***) | |
| 227 | ||
| 7341 | 228 | (*Because A and B could differ on the "other" part of the state, | 
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
8041diff
changeset | 229 | cannot generalize to | 
| 7341 | 230 | project_set h (A Int B) = project_set h A Int project_set h B | 
| 231 | *) | |
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 232 | Goal "project_set h ((extend_set h A) Int B) = A Int (project_set h B)"; | 
| 7341 | 233 | by Auto_tac; | 
| 234 | qed "project_set_extend_set_Int"; | |
| 235 | ||
| 9609 | 236 | (*Unused, but interesting?*) | 
| 237 | Goal "project_set h ((extend_set h A) Un B) = A Un (project_set h B)"; | |
| 238 | by Auto_tac; | |
| 239 | qed "project_set_extend_set_Un"; | |
| 240 | ||
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 241 | Goal "project_set h (A Int B) <= (project_set h A) Int (project_set h B)"; | 
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
8041diff
changeset | 242 | by Auto_tac; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
8041diff
changeset | 243 | qed "project_set_Int_subset"; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
8041diff
changeset | 244 | |
| 7341 | 245 | Goal "extend_set h (A Un B) = extend_set h A Un extend_set h B"; | 
| 6297 | 246 | by Auto_tac; | 
| 247 | qed "extend_set_Un_distrib"; | |
| 248 | ||
| 7341 | 249 | Goal "extend_set h (A Int B) = extend_set h A Int extend_set h B"; | 
| 6297 | 250 | by Auto_tac; | 
| 251 | qed "extend_set_Int_distrib"; | |
| 252 | ||
| 7341 | 253 | Goal "extend_set h (INTER A B) = (INT x:A. extend_set h (B x))"; | 
| 254 | by Auto_tac; | |
| 6834 
44da4a2a9ef3
renamed UNION_... to UN_..., INTER_... to INT_... (to fit the convention)
 paulson parents: 
6822diff
changeset | 255 | qed "extend_set_INT_distrib"; | 
| 6647 | 256 | |
| 7341 | 257 | Goal "extend_set h (A - B) = extend_set h A - extend_set h B"; | 
| 6297 | 258 | by Auto_tac; | 
| 259 | qed "extend_set_Diff_distrib"; | |
| 260 | ||
| 7341 | 261 | Goal "extend_set h (Union A) = (UN X:A. extend_set h X)"; | 
| 6297 | 262 | by (Blast_tac 1); | 
| 263 | qed "extend_set_Union"; | |
| 264 | ||
| 7341 | 265 | Goalw [extend_set_def] "(extend_set h A <= - extend_set h B) = (A <= - B)"; | 
| 6297 | 266 | by Auto_tac; | 
| 267 | qed "extend_set_subset_Compl_eq"; | |
| 268 | ||
| 7341 | 269 | |
| 6297 | 270 | (*** extend_act ***) | 
| 271 | ||
| 8216 | 272 | (*Can't strengthen it to | 
| 273 | ((h(s,y), h(s',y')) : extend_act h act) = ((s, s') : act & y=y') | |
| 274 | because h doesn't have to be injective in the 2nd argument*) | |
| 6297 | 275 | Goalw [extend_act_def] | 
| 276 | "((h(s,y), h(s',y)) : extend_act h act) = ((s, s') : act)"; | |
| 277 | by Auto_tac; | |
| 278 | qed "mem_extend_act_iff"; | |
| 279 | AddIffs [mem_extend_act_iff]; | |
| 280 | ||
| 7878 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 281 | (*Converse fails: (z,z') would include actions that changed the g-part*) | 
| 7341 | 282 | Goalw [extend_act_def] | 
| 283 | "(z, z') : extend_act h act ==> (f z, f z') : act"; | |
| 284 | by Auto_tac; | |
| 285 | qed "extend_act_D"; | |
| 286 | ||
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 287 | Goalw [extend_act_def, project_act_def] | 
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
8041diff
changeset | 288 | "project_act h (extend_act h act) = act"; | 
| 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
8041diff
changeset | 289 | by (Blast_tac 1); | 
| 9337 
58bd51302b21
used bounded quantification in definition of guarantees and other minor
 paulson parents: 
9190diff
changeset | 290 | qed "extend_act_inverse"; | 
| 8055 
bb15396278fb
abolition of localTo: instead "guarantees" has local vars as extra argument
 paulson parents: 
8041diff
changeset | 291 | |
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 292 | Goalw [extend_act_def, project_act_def] | 
| 7878 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 293 | "project_act h (Restrict C (extend_act h act)) = \ | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 294 | \ Restrict (project_set h C) act"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 295 | by (Blast_tac 1); | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 296 | qed "project_act_extend_act_restrict"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 297 | Addsimps [project_act_extend_act_restrict]; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 298 | |
| 7826 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 299 | Goalw [extend_act_def, project_act_def] | 
| 7878 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 300 | "act' <= extend_act h act ==> project_act h act' <= act"; | 
| 7826 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 301 | by (Force_tac 1); | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 302 | qed "subset_extend_act_D"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 303 | |
| 6297 | 304 | Goal "inj (extend_act h)"; | 
| 7341 | 305 | by (rtac inj_on_inverseI 1); | 
| 306 | by (rtac extend_act_inverse 1); | |
| 6297 | 307 | qed "inj_extend_act"; | 
| 308 | ||
| 309 | Goalw [extend_set_def, extend_act_def] | |
| 310 | "extend_act h act ^^ (extend_set h A) = extend_set h (act ^^ A)"; | |
| 311 | by (Force_tac 1); | |
| 312 | qed "extend_act_Image"; | |
| 313 | Addsimps [extend_act_Image]; | |
| 314 | ||
| 7826 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 315 | Goalw [extend_act_def] "(extend_act h act' <= extend_act h act) = (act'<=act)"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 316 | by Auto_tac; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 317 | qed "extend_act_strict_mono"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 318 | |
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 319 | AddIffs [extend_act_strict_mono, inj_extend_act RS inj_eq]; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 320 | (*The latter theorem is (extend_act h act' = extend_act h act) = (act'=act) *) | 
| 6297 | 321 | |
| 322 | Goalw [extend_set_def, extend_act_def] | |
| 323 | "Domain (extend_act h act) = extend_set h (Domain act)"; | |
| 324 | by (Force_tac 1); | |
| 325 | qed "Domain_extend_act"; | |
| 326 | ||
| 7341 | 327 | Goalw [extend_act_def] | 
| 6297 | 328 | "extend_act h Id = Id"; | 
| 329 | by (force_tac (claset() addIs [h_f_g_eq RS sym], simpset()) 1); | |
| 330 | qed "extend_act_Id"; | |
| 7341 | 331 | |
| 332 | Goalw [project_act_def] | |
| 7878 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 333 | "!!z z'. (z, z') : act ==> (f z, f z') : project_act h act"; | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 334 | by (force_tac (claset(), | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 335 | simpset() addsimps [split_extended_all]) 1); | 
| 7341 | 336 | qed "project_act_I"; | 
| 337 | ||
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 338 | Goalw [project_act_def] "project_act h Id = Id"; | 
| 7341 | 339 | by (Force_tac 1); | 
| 340 | qed "project_act_Id"; | |
| 341 | ||
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 342 | Goalw [project_act_def] | 
| 7915 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7880diff
changeset | 343 | "Domain (project_act h act) = project_set h (Domain act)"; | 
| 7878 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 344 | by (force_tac (claset(), | 
| 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 345 | simpset() addsimps [split_extended_all]) 1); | 
| 7482 | 346 | qed "Domain_project_act"; | 
| 347 | ||
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 348 | Addsimps [extend_act_Id]; | 
| 6297 | 349 | |
| 350 | ||
| 351 | (**** extend ****) | |
| 352 | ||
| 353 | (*** Basic properties ***) | |
| 354 | ||
| 7341 | 355 | Goalw [extend_def] "Init (extend h F) = extend_set h (Init F)"; | 
| 6297 | 356 | by Auto_tac; | 
| 357 | qed "Init_extend"; | |
| 358 | ||
| 7880 
62fb24e28e5e
exchanged the first two args of "project" and "drop_prog"
 paulson parents: 
7878diff
changeset | 359 | Goalw [project_def] "Init (project h C F) = project_set h (Init F)"; | 
| 7341 | 360 | by Auto_tac; | 
| 361 | qed "Init_project"; | |
| 362 | ||
| 6297 | 363 | Goal "Acts (extend h F) = (extend_act h `` Acts F)"; | 
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 364 | by (simp_tac (simpset() addsimps [extend_def, insert_Id_image_Acts]) 1); | 
| 6297 | 365 | qed "Acts_extend"; | 
| 366 | ||
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 367 | Goal "Acts(project h C F) = insert Id (project_act h `` Restrict C `` Acts F)"; | 
| 7546 
36b26759147e
project_act no longer has a special case to allow identity actions
 paulson parents: 
7538diff
changeset | 368 | by (auto_tac (claset(), | 
| 7341 | 369 | simpset() addsimps [project_def, image_iff])); | 
| 370 | qed "Acts_project"; | |
| 371 | ||
| 372 | Addsimps [Init_extend, Init_project, Acts_extend, Acts_project]; | |
| 6297 | 373 | |
| 374 | Goalw [SKIP_def] "extend h SKIP = SKIP"; | |
| 375 | by (rtac program_equalityI 1); | |
| 7341 | 376 | by Auto_tac; | 
| 377 | qed "extend_SKIP"; | |
| 8122 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 378 | Addsimps [export extend_SKIP]; | 
| 7341 | 379 | |
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 380 | Goal "project_set h UNIV = UNIV"; | 
| 7537 | 381 | by Auto_tac; | 
| 382 | qed "project_set_UNIV"; | |
| 7878 
43b03d412b82
working version with localTo[C] instead of localTo
 paulson parents: 
7826diff
changeset | 383 | Addsimps [project_set_UNIV]; | 
| 7537 | 384 | |
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 385 | Goal "project_set h (Union A) = (UN X:A. project_set h X)"; | 
| 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 386 | by (Blast_tac 1); | 
| 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 387 | qed "project_set_Union"; | 
| 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 388 | |
| 7915 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7880diff
changeset | 389 | Goal "project h C (extend h F) = \ | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7880diff
changeset | 390 | \ mk_program (Init F, Restrict (project_set h C) `` Acts F)"; | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7880diff
changeset | 391 | by (rtac program_equalityI 1); | 
| 9403 
aad13b59b8d9
much tidying in connection with the 2nd UNITY paper
 paulson parents: 
9345diff
changeset | 392 | by (asm_simp_tac (simpset() addsimps [image_eq_UN]) 2); | 
| 7915 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7880diff
changeset | 393 | by (Simp_tac 1); | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7880diff
changeset | 394 | qed "project_extend_eq"; | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7880diff
changeset | 395 | |
| 7537 | 396 | (*ALL act: Acts F. Domain act is MUCH TOO STRONG since Domain Id = UNIV!*) | 
| 397 | Goal "UNIV <= project_set h C \ | |
| 7880 
62fb24e28e5e
exchanged the first two args of "project" and "drop_prog"
 paulson parents: 
7878diff
changeset | 398 | \ ==> project h C (extend h F) = F"; | 
| 7915 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7880diff
changeset | 399 | by (asm_simp_tac (simpset() addsimps [project_extend_eq, image_eq_UN, | 
| 
c7fd7eb3b0ef
ALMOST working version: LocalTo results commented out
 paulson parents: 
7880diff
changeset | 400 | subset_UNIV RS subset_trans RS Restrict_triv]) 1); | 
| 7341 | 401 | qed "extend_inverse"; | 
| 402 | Addsimps [extend_inverse]; | |
| 6297 | 403 | |
| 404 | Goal "inj (extend h)"; | |
| 7341 | 405 | by (rtac inj_on_inverseI 1); | 
| 406 | by (rtac extend_inverse 1); | |
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 407 | by (Force_tac 1); | 
| 6297 | 408 | qed "inj_extend"; | 
| 409 | ||
| 410 | Goal "extend h (F Join G) = extend h F Join extend h G"; | |
| 411 | by (rtac program_equalityI 1); | |
| 7537 | 412 | by (simp_tac (simpset() addsimps [image_Un]) 2); | 
| 6297 | 413 | by (simp_tac (simpset() addsimps [extend_set_Int_distrib]) 1); | 
| 414 | qed "extend_Join"; | |
| 415 | Addsimps [extend_Join]; | |
| 416 | ||
| 6647 | 417 | Goal "extend h (JOIN I F) = (JN i:I. extend h (F i))"; | 
| 418 | by (rtac program_equalityI 1); | |
| 7537 | 419 | by (simp_tac (simpset() addsimps [image_UN]) 2); | 
| 6834 
44da4a2a9ef3
renamed UNION_... to UN_..., INTER_... to INT_... (to fit the convention)
 paulson parents: 
6822diff
changeset | 420 | by (simp_tac (simpset() addsimps [extend_set_INT_distrib]) 1); | 
| 6647 | 421 | qed "extend_JN"; | 
| 422 | Addsimps [extend_JN]; | |
| 423 | ||
| 7689 | 424 | (** These monotonicity results look natural but are UNUSED **) | 
| 425 | ||
| 426 | Goal "F <= G ==> extend h F <= extend h G"; | |
| 427 | by (full_simp_tac (simpset() addsimps [component_eq_subset]) 1); | |
| 428 | by Auto_tac; | |
| 429 | qed "extend_mono"; | |
| 430 | ||
| 7880 
62fb24e28e5e
exchanged the first two args of "project" and "drop_prog"
 paulson parents: 
7878diff
changeset | 431 | Goal "F <= G ==> project h C F <= project h C G"; | 
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 432 | by (full_simp_tac (simpset() addsimps [component_eq_subset]) 1); | 
| 7689 | 433 | by Auto_tac; | 
| 434 | qed "project_mono"; | |
| 435 | ||
| 436 | ||
| 6536 | 437 | (*** Safety: co, stable ***) | 
| 6297 | 438 | |
| 6536 | 439 | Goal "(extend h F : (extend_set h A) co (extend_set h B)) = \ | 
| 440 | \ (F : A co B)"; | |
| 6297 | 441 | by (simp_tac (simpset() addsimps [constrains_def]) 1); | 
| 442 | qed "extend_constrains"; | |
| 443 | ||
| 444 | Goal "(extend h F : stable (extend_set h A)) = (F : stable A)"; | |
| 445 | by (asm_simp_tac (simpset() addsimps [stable_def, extend_constrains]) 1); | |
| 446 | qed "extend_stable"; | |
| 447 | ||
| 448 | Goal "(extend h F : invariant (extend_set h A)) = (F : invariant A)"; | |
| 449 | by (asm_simp_tac (simpset() addsimps [invariant_def, extend_stable]) 1); | |
| 450 | qed "extend_invariant"; | |
| 451 | ||
| 9337 
58bd51302b21
used bounded quantification in definition of guarantees and other minor
 paulson parents: 
9190diff
changeset | 452 | (*Projects the state predicates in the property satisfied by extend h F. | 
| 
58bd51302b21
used bounded quantification in definition of guarantees and other minor
 paulson parents: 
9190diff
changeset | 453 | Converse fails: A and B may differ in their extra variables*) | 
| 7826 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 454 | Goal "extend h F : A co B ==> F : (project_set h A) co (project_set h B)"; | 
| 8069 
19b9f92ca503
working with weak LeadsTo in guarantees precondition\!
 paulson parents: 
8055diff
changeset | 455 | by (auto_tac (claset(), simpset() addsimps [constrains_def])); | 
| 7826 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 456 | by (Force_tac 1); | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 457 | qed "extend_constrains_project_set"; | 
| 
c6a8b73b6c2a
working shapshot with "projecting" and "extending"
 paulson parents: 
7689diff
changeset | 458 | |
| 8041 | 459 | Goal "extend h F : stable A ==> F : stable (project_set h A)"; | 
| 460 | by (asm_full_simp_tac | |
| 461 | (simpset() addsimps [stable_def, extend_constrains_project_set]) 1); | |
| 462 | qed "extend_stable_project_set"; | |
| 463 | ||
| 7341 | 464 | |
| 465 | (*** Weak safety primitives: Co, Stable ***) | |
| 6297 | 466 | |
| 467 | Goal "p : reachable (extend h F) ==> f p : reachable F"; | |
| 468 | by (etac reachable.induct 1); | |
| 469 | by (auto_tac | |
| 470 | (claset() addIs reachable.intrs, | |
| 7341 | 471 | simpset() addsimps [extend_act_def, image_iff])); | 
| 6297 | 472 | qed "reachable_extend_f"; | 
| 473 | ||
| 474 | Goal "h(s,y) : reachable (extend h F) ==> s : reachable F"; | |
| 475 | by (force_tac (claset() addSDs [reachable_extend_f], simpset()) 1); | |
| 476 | qed "h_reachable_extend"; | |
| 477 | ||
| 478 | Goalw [extend_set_def] | |
| 479 | "reachable (extend h F) = extend_set h (reachable F)"; | |
| 480 | by (rtac equalityI 1); | |
| 481 | by (force_tac (claset() addIs [h_f_g_eq RS sym] | |
| 482 | addSDs [reachable_extend_f], | |
| 483 | simpset()) 1); | |
| 484 | by (Clarify_tac 1); | |
| 485 | by (etac reachable.induct 1); | |
| 486 | by (ALLGOALS (force_tac (claset() addIs reachable.intrs, | |
| 487 | simpset()))); | |
| 488 | qed "reachable_extend_eq"; | |
| 489 | ||
| 6536 | 490 | Goal "(extend h F : (extend_set h A) Co (extend_set h B)) = \ | 
| 491 | \ (F : A Co B)"; | |
| 6297 | 492 | by (simp_tac | 
| 493 | (simpset() addsimps [Constrains_def, reachable_extend_eq, | |
| 494 | extend_constrains, extend_set_Int_distrib RS sym]) 1); | |
| 495 | qed "extend_Constrains"; | |
| 496 | ||
| 497 | Goal "(extend h F : Stable (extend_set h A)) = (F : Stable A)"; | |
| 498 | by (simp_tac (simpset() addsimps [Stable_def, extend_Constrains]) 1); | |
| 499 | qed "extend_Stable"; | |
| 500 | ||
| 6647 | 501 | Goal "(extend h F : Always (extend_set h A)) = (F : Always A)"; | 
| 502 | by (asm_simp_tac (simpset() addsimps [Always_def, extend_Stable]) 1); | |
| 503 | qed "extend_Always"; | |
| 504 | ||
| 6297 | 505 | |
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 506 | (** Safety and "project" **) | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 507 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 508 | (** projection: monotonicity for safety **) | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 509 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 510 | Goal "D <= C ==> \ | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 511 | \ project_act h (Restrict D act) <= project_act h (Restrict C act)"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 512 | by (auto_tac (claset(), simpset() addsimps [project_act_def])); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 513 | qed "project_act_mono"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 514 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 515 | Goal "[| D <= C; project h C F : A co B |] ==> project h D F : A co B"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 516 | by (auto_tac (claset(), simpset() addsimps [constrains_def])); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 517 | by (dtac project_act_mono 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 518 | by (Blast_tac 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 519 | qed "project_constrains_mono"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 520 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 521 | Goal "[| D <= C; project h C F : stable A |] ==> project h D F : stable A"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 522 | by (asm_full_simp_tac | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 523 | (simpset() addsimps [stable_def, project_constrains_mono]) 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 524 | qed "project_stable_mono"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 525 | |
| 9337 
58bd51302b21
used bounded quantification in definition of guarantees and other minor
 paulson parents: 
9190diff
changeset | 526 | (*Key lemma used in several proofs about project and co*) | 
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 527 | Goalw [constrains_def] | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 528 | "(project h C F : A co B) = \ | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 529 | \ (F : (C Int extend_set h A) co (extend_set h B) & A <= B)"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 530 | by (auto_tac (claset() addSIs [project_act_I], simpset() addsimps [ball_Un])); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 531 | by (force_tac (claset() addSIs [project_act_I] addSDs [subsetD], simpset()) 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 532 | (*the <== direction*) | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 533 | by (rewtac project_act_def); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 534 | by (force_tac (claset() addSDs [subsetD], simpset()) 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 535 | qed "project_constrains"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 536 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 537 | Goalw [stable_def] | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 538 | "(project h UNIV F : stable A) = (F : stable (extend_set h A))"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 539 | by (simp_tac (simpset() addsimps [project_constrains]) 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 540 | qed "project_stable"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 541 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 542 | Goal "F : stable (extend_set h A) ==> project h C F : stable A"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 543 | by (dtac (project_stable RS iffD2) 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 544 | by (blast_tac (claset() addIs [project_stable_mono]) 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 545 | qed "project_stable_I"; | 
| 8122 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 546 | |
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 547 | Goal "A Int extend_set h ((project_set h A) Int B) = A Int extend_set h B"; | 
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 548 | by (auto_tac (claset(), simpset() addsimps [split_extended_all])); | 
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 549 | qed "Int_extend_set_lemma"; | 
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 550 | |
| 9337 
58bd51302b21
used bounded quantification in definition of guarantees and other minor
 paulson parents: 
9190diff
changeset | 551 | (*Strange (look at occurrences of C) but used in leadsETo proofs*) | 
| 8122 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 552 | Goal "G : C co B ==> project h C G : project_set h C co project_set h B"; | 
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 553 | by (full_simp_tac (simpset() addsimps [constrains_def, project_def, | 
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 554 | project_act_def]) 1); | 
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 555 | by (Blast_tac 1); | 
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 556 | qed "project_constrains_project_set"; | 
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 557 | |
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 558 | Goal "G : stable C ==> project h C G : stable (project_set h C)"; | 
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 559 | by (asm_full_simp_tac (simpset() addsimps [stable_def, | 
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 560 | project_constrains_project_set]) 1); | 
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 561 | qed "project_stable_project_set"; | 
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 562 | |
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8111diff
changeset | 563 | |
| 6297 | 564 | (*** Progress: transient, ensures ***) | 
| 565 | ||
| 566 | Goal "(extend h F : transient (extend_set h A)) = (F : transient A)"; | |
| 567 | by (auto_tac (claset(), | |
| 568 | simpset() addsimps [transient_def, extend_set_subset_Compl_eq, | |
| 569 | Domain_extend_act])); | |
| 570 | qed "extend_transient"; | |
| 571 | ||
| 6536 | 572 | Goal "(extend h F : (extend_set h A) ensures (extend_set h B)) = \ | 
| 573 | \ (F : A ensures B)"; | |
| 6297 | 574 | by (simp_tac | 
| 575 | (simpset() addsimps [ensures_def, extend_constrains, extend_transient, | |
| 576 | extend_set_Un_distrib RS sym, | |
| 577 | extend_set_Diff_distrib RS sym]) 1); | |
| 578 | qed "extend_ensures"; | |
| 579 | ||
| 6536 | 580 | Goal "F : A leadsTo B \ | 
| 581 | \ ==> extend h F : (extend_set h A) leadsTo (extend_set h B)"; | |
| 6297 | 582 | by (etac leadsTo_induct 1); | 
| 583 | by (asm_simp_tac (simpset() addsimps [leadsTo_UN, extend_set_Union]) 3); | |
| 584 | by (blast_tac (claset() addIs [leadsTo_Trans]) 2); | |
| 585 | by (asm_simp_tac (simpset() addsimps [leadsTo_Basis, extend_ensures]) 1); | |
| 586 | qed "leadsTo_imp_extend_leadsTo"; | |
| 587 | ||
| 588 | (*** Proving the converse takes some doing! ***) | |
| 589 | ||
| 8111 | 590 | Goal "(x : slice C y) = (h(x,y) : C)"; | 
| 591 | by (simp_tac (simpset() addsimps [slice_def]) 1); | |
| 592 | qed "slice_iff"; | |
| 593 | ||
| 594 | AddIffs [slice_iff]; | |
| 595 | ||
| 596 | Goal "slice (Union S) y = (UN x:S. slice x y)"; | |
| 6297 | 597 | by Auto_tac; | 
| 598 | qed "slice_Union"; | |
| 599 | ||
| 8111 | 600 | Goal "slice (extend_set h A) y = A"; | 
| 6297 | 601 | by Auto_tac; | 
| 602 | qed "slice_extend_set"; | |
| 603 | ||
| 8111 | 604 | Goal "project_set h A = (UN y. slice A y)"; | 
| 6297 | 605 | by Auto_tac; | 
| 7482 | 606 | qed "project_set_is_UN_slice"; | 
| 6297 | 607 | |
| 8111 | 608 | Goalw [transient_def] "extend h F : transient A ==> F : transient (slice A y)"; | 
| 6297 | 609 | by Auto_tac; | 
| 610 | by (rtac bexI 1); | |
| 611 | by Auto_tac; | |
| 612 | by (force_tac (claset(), simpset() addsimps [extend_act_def]) 1); | |
| 613 | qed "extend_transient_slice"; | |
| 614 | ||
| 8111 | 615 | (*Converse?*) | 
| 616 | Goal "extend h F : A co B ==> F : (slice A y) co (slice B y)"; | |
| 617 | by (auto_tac (claset(), simpset() addsimps [constrains_def])); | |
| 618 | qed "extend_constrains_slice"; | |
| 619 | ||
| 7482 | 620 | Goal "extend h F : A ensures B ==> F : (slice A y) ensures (project_set h B)"; | 
| 8111 | 621 | by (auto_tac (claset(), | 
| 622 | simpset() addsimps [ensures_def, extend_constrains, | |
| 623 | extend_transient])); | |
| 624 | by (etac (extend_transient_slice RS transient_strengthen) 2); | |
| 625 | by (etac (extend_constrains_slice RS constrains_weaken) 1); | |
| 626 | by Auto_tac; | |
| 6297 | 627 | qed "extend_ensures_slice"; | 
| 628 | ||
| 7482 | 629 | Goal "ALL y. F : (slice B y) leadsTo CU ==> F : (project_set h B) leadsTo CU"; | 
| 630 | by (simp_tac (simpset() addsimps [project_set_is_UN_slice]) 1); | |
| 6297 | 631 | by (blast_tac (claset() addIs [leadsTo_UN]) 1); | 
| 8111 | 632 | qed "leadsTo_slice_project_set"; | 
| 6297 | 633 | |
| 6536 | 634 | Goal "extend h F : AU leadsTo BU \ | 
| 7482 | 635 | \ ==> ALL y. F : (slice AU y) leadsTo (project_set h BU)"; | 
| 6297 | 636 | by (etac leadsTo_induct 1); | 
| 8111 | 637 | by (asm_simp_tac (simpset() addsimps [leadsTo_UN, slice_Union]) 3); | 
| 638 | by (blast_tac (claset() addIs [leadsTo_slice_project_set, leadsTo_Trans]) 2); | |
| 6297 | 639 | by (blast_tac (claset() addIs [extend_ensures_slice, leadsTo_Basis]) 1); | 
| 640 | qed_spec_mp "extend_leadsTo_slice"; | |
| 641 | ||
| 6536 | 642 | Goal "(extend h F : (extend_set h A) leadsTo (extend_set h B)) = \ | 
| 643 | \ (F : A leadsTo B)"; | |
| 6297 | 644 | by Safe_tac; | 
| 645 | by (etac leadsTo_imp_extend_leadsTo 2); | |
| 646 | by (dtac extend_leadsTo_slice 1); | |
| 647 | by (full_simp_tac (simpset() addsimps [slice_extend_set]) 1); | |
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 648 | qed "extend_leadsTo"; | 
| 6297 | 649 | |
| 6536 | 650 | Goal "(extend h F : (extend_set h A) LeadsTo (extend_set h B)) = \ | 
| 651 | \ (F : A LeadsTo B)"; | |
| 6454 | 652 | by (simp_tac | 
| 653 | (simpset() addsimps [LeadsTo_def, reachable_extend_eq, | |
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 654 | extend_leadsTo, extend_set_Int_distrib RS sym]) 1); | 
| 6454 | 655 | qed "extend_LeadsTo"; | 
| 656 | ||
| 6297 | 657 | |
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 658 | (*** preserves ***) | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 659 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 660 | Goal "G : preserves (v o f) ==> project h C G : preserves v"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 661 | by (auto_tac (claset(), | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 662 | simpset() addsimps [preserves_def, project_stable_I, | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 663 | extend_set_eq_Collect])); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 664 | qed "project_preserves_I"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 665 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 666 | (*to preserve f is to preserve the whole original state*) | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 667 | Goal "G : preserves f ==> project h C G : preserves id"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 668 | by (asm_simp_tac (simpset() addsimps [project_preserves_I]) 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 669 | qed "project_preserves_id_I"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 670 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 671 | Goal "(extend h G : preserves (v o f)) = (G : preserves v)"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 672 | by (auto_tac (claset(), | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 673 | simpset() addsimps [preserves_def, extend_stable RS sym, | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 674 | extend_set_eq_Collect])); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 675 | qed "extend_preserves"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 676 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 677 | Goal "inj h ==> (extend h G : preserves g)"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 678 | by (auto_tac (claset(), | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 679 | simpset() addsimps [preserves_def, extend_def, extend_act_def, | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 680 | stable_def, constrains_def, g_def])); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 681 | qed "inj_extend_preserves"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 682 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 683 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 684 | (*** Guarantees ***) | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 685 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 686 | Goal "project h UNIV ((extend h F) Join G) = F Join (project h UNIV G)"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 687 | by (rtac program_equalityI 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 688 | by (asm_simp_tac | 
| 9337 
58bd51302b21
used bounded quantification in definition of guarantees and other minor
 paulson parents: 
9190diff
changeset | 689 | (simpset() addsimps [image_eq_UN, UN_Un, extend_act_inverse]) 2); | 
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 690 | by (simp_tac (simpset() addsimps [project_set_extend_set_Int]) 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 691 | qed "project_extend_Join"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 692 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 693 | Goal "(extend h F) Join G = extend h H ==> H = F Join (project h UNIV G)"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 694 | by (dres_inst_tac [("f", "project h UNIV")] arg_cong 1);
 | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 695 | by (asm_full_simp_tac (simpset() addsimps [project_extend_Join]) 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 696 | qed "extend_Join_eq_extend_D"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 697 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 698 | (** Strong precondition and postcondition; only useful when | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 699 | the old and new state sets are in bijection **) | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 700 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 701 | Goal "F : X guarantees[v] Y ==> \ | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 702 | \ extend h F : (extend h `` X) guarantees[v o f] (extend h `` Y)"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 703 | by (rtac guaranteesI 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 704 | by Auto_tac; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 705 | by (blast_tac (claset() addIs [project_preserves_I] | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 706 | addDs [extend_Join_eq_extend_D, guaranteesD]) 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 707 | qed "guarantees_imp_extend_guarantees"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 708 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 709 | Goal "extend h F : (extend h `` X) guarantees[v o f] (extend h `` Y) \ | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 710 | \ ==> F : X guarantees[v] Y"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 711 | by (auto_tac (claset(), simpset() addsimps [guar_def])); | 
| 9337 
58bd51302b21
used bounded quantification in definition of guarantees and other minor
 paulson parents: 
9190diff
changeset | 712 | by (dres_inst_tac [("x", "extend h G")] bspec 1);
 | 
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 713 | by (asm_full_simp_tac | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 714 | (simpset() delsimps [extend_Join] | 
| 9337 
58bd51302b21
used bounded quantification in definition of guarantees and other minor
 paulson parents: 
9190diff
changeset | 715 | addsimps [extend_Join RS sym, inj_extend RS inj_image_mem_iff]) 2); | 
| 
58bd51302b21
used bounded quantification in definition of guarantees and other minor
 paulson parents: 
9190diff
changeset | 716 | by (asm_simp_tac (simpset() addsimps [extend_preserves]) 1); | 
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 717 | qed "extend_guarantees_imp_guarantees"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 718 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 719 | Goal "(extend h F : (extend h `` X) guarantees[v o f] (extend h `` Y)) = \ | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 720 | \ (F : X guarantees[v] Y)"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 721 | by (blast_tac (claset() addIs [guarantees_imp_extend_guarantees, | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 722 | extend_guarantees_imp_guarantees]) 1); | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 723 | qed "extend_guarantees_eq"; | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 724 | |
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8216diff
changeset | 725 | |
| 6297 | 726 | Close_locale "Extend"; | 
| 7482 | 727 | |
| 728 | (*Close_locale should do this! | |
| 9337 
58bd51302b21
used bounded quantification in definition of guarantees and other minor
 paulson parents: 
9190diff
changeset | 729 | Delsimps [f_h_eq, extend_set_inverse, f_image_extend_set, extend_act_Image]; | 
| 7482 | 730 | Delrules [make_elim h_inject1]; | 
| 731 | *) |