| author | wenzelm | 
| Tue, 30 Dec 2008 21:49:09 +0100 | |
| changeset 29261 | 7ee78cc8ef2c | 
| parent 24893 | b8ef7afe3a6b | 
| child 32960 | 69916a850301 | 
| permissions | -rw-r--r-- | 
| 11479 | 1 | (* Title: ZF/UNITY/Comp.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Sidi O Ehmety, Computer Laboratory | |
| 4 | Copyright 1998 University of Cambridge | |
| 5 | ||
| 6 | From Chandy and Sanders, "Reasoning About Program Composition", | |
| 7 | Technical Report 2000-003, University of Florida, 2000. | |
| 8 | ||
| 9 | Revised by Sidi Ehmety on January 2001 | |
| 10 | ||
| 11 | Added: a strong form of the order relation over component and localize | |
| 12 | ||
| 13 | Theory ported from HOL. | |
| 14 | ||
| 15 | *) | |
| 16 | ||
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 17 | header{*Composition*}
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 18 | |
| 16417 | 19 | theory Comp imports Union Increasing begin | 
| 11479 | 20 | |
| 24893 | 21 | definition | 
| 22 | component :: "[i,i]=>o" (infixl "component" 65) where | |
| 11479 | 23 | "F component H == (EX G. F Join G = H)" | 
| 24 | ||
| 24893 | 25 | definition | 
| 26 | strict_component :: "[i,i]=>o" (infixl "strict'_component" 65) where | |
| 11479 | 27 | "F strict_component H == F component H & F~=H" | 
| 28 | ||
| 24893 | 29 | definition | 
| 11479 | 30 | (* A stronger form of the component relation *) | 
| 24893 | 31 | component_of :: "[i,i]=>o" (infixl "component'_of" 65) where | 
| 11479 | 32 | "F component_of H == (EX G. F ok G & F Join G = H)" | 
| 33 | ||
| 24893 | 34 | definition | 
| 35 | strict_component_of :: "[i,i]=>o" (infixl "strict'_component'_of" 65) where | |
| 11479 | 36 | "F strict_component_of H == F component_of H & F~=H" | 
| 37 | ||
| 24893 | 38 | definition | 
| 12195 | 39 | (* Preserves a state function f, in particular a variable *) | 
| 24893 | 40 | preserves :: "(i=>i)=>i" where | 
| 11479 | 41 |   "preserves(f) == {F:program. ALL z. F: stable({s:state. f(s) = z})}"
 | 
| 42 | ||
| 24893 | 43 | definition | 
| 44 | fun_pair :: "[i=>i, i =>i] =>(i=>i)" where | |
| 14046 | 45 | "fun_pair(f, g) == %x. <f(x), g(x)>" | 
| 46 | ||
| 24893 | 47 | definition | 
| 48 | localize :: "[i=>i, i] => i" where | |
| 11479 | 49 | "localize(f,F) == mk_program(Init(F), Acts(F), | 
| 14046 | 50 | AllowedActs(F) Int (UN G:preserves(f). Acts(G)))" | 
| 11479 | 51 | |
| 52 | ||
| 14093 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 53 | (*** component and strict_component relations ***) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 54 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 55 | lemma componentI: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 56 | "H component F | H component G ==> H component (F Join G)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 57 | apply (unfold component_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 58 | apply (rule_tac x = "Ga Join G" in exI) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 59 | apply (rule_tac [2] x = "G Join F" in exI) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 60 | apply (auto simp add: Join_ac) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 61 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 62 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 63 | lemma component_eq_subset: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 64 | "G \<in> program ==> (F component G) <-> | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 65 | (Init(G) <= Init(F) & Acts(F) <= Acts(G) & AllowedActs(G) <= AllowedActs(F))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 66 | apply (unfold component_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 67 | apply (rule exI) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 68 | apply (rule program_equalityI, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 69 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 70 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 71 | lemma component_SKIP [simp]: "F \<in> program ==> SKIP component F" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 72 | apply (unfold component_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 73 | apply (rule_tac x = F in exI) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 74 | apply (force intro: Join_SKIP_left) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 75 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 76 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 77 | lemma component_refl [simp]: "F \<in> program ==> F component F" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 78 | apply (unfold component_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 79 | apply (rule_tac x = F in exI) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 80 | apply (force intro: Join_SKIP_right) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 81 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 82 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 83 | lemma SKIP_minimal: "F component SKIP ==> programify(F) = SKIP" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 84 | apply (rule program_equalityI) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 85 | apply (simp_all add: component_eq_subset, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 86 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 87 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 88 | lemma component_Join1: "F component (F Join G)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 89 | by (unfold component_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 90 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 91 | lemma component_Join2: "G component (F Join G)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 92 | apply (unfold component_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 93 | apply (simp (no_asm) add: Join_commute) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 94 | apply blast | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 95 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 96 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 97 | lemma Join_absorb1: "F component G ==> F Join G = G" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 98 | by (auto simp add: component_def Join_left_absorb) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 99 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 100 | lemma Join_absorb2: "G component F ==> F Join G = F" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 101 | by (auto simp add: Join_ac component_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 102 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 103 | lemma JN_component_iff: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 104 | "H \<in> program==>(JOIN(I,F) component H) <-> (\<forall>i \<in> I. F(i) component H)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 105 | apply (case_tac "I=0", force) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 106 | apply (simp (no_asm_simp) add: component_eq_subset) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 107 | apply auto | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 108 | apply blast | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 109 | apply (rename_tac "y") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 110 | apply (drule_tac c = y and A = "AllowedActs (H)" in subsetD) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 111 | apply (blast elim!: not_emptyE)+ | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 112 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 113 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 114 | lemma component_JN: "i \<in> I ==> F(i) component (\<Squnion>i \<in> I. (F(i)))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 115 | apply (unfold component_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 116 | apply (blast intro: JN_absorb) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 117 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 118 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 119 | lemma component_trans: "[| F component G; G component H |] ==> F component H" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 120 | apply (unfold component_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 121 | apply (blast intro: Join_assoc [symmetric]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 122 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 123 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 124 | lemma component_antisym: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 125 | "[| F \<in> program; G \<in> program |] ==>(F component G & G component F) --> F = G" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 126 | apply (simp (no_asm_simp) add: component_eq_subset) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 127 | apply clarify | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 128 | apply (rule program_equalityI, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 129 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 130 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 131 | lemma Join_component_iff: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 132 | "H \<in> program ==> | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 133 | ((F Join G) component H) <-> (F component H & G component H)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 134 | apply (simp (no_asm_simp) add: component_eq_subset) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 135 | apply blast | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 136 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 137 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 138 | lemma component_constrains: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 139 | "[| F component G; G \<in> A co B; F \<in> program |] ==> F \<in> A co B" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 140 | apply (frule constrainsD2) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 141 | apply (auto simp add: constrains_def component_eq_subset) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 142 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 143 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 144 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 145 | (*** preserves ***) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 146 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 147 | lemma preserves_is_safety_prop [simp]: "safety_prop(preserves(f))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 148 | apply (unfold preserves_def safety_prop_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 149 | apply (auto dest: ActsD simp add: stable_def constrains_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 150 | apply (drule_tac c = act in subsetD, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 151 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 152 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 153 | lemma preservesI [rule_format]: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 154 |      "\<forall>z. F \<in> stable({s \<in> state. f(s) = z})  ==> F \<in> preserves(f)"
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 155 | apply (auto simp add: preserves_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 156 | apply (blast dest: stableD2) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 157 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 158 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 159 | lemma preserves_imp_eq: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 160 | "[| F \<in> preserves(f); act \<in> Acts(F); <s,t> \<in> act |] ==> f(s) = f(t)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 161 | apply (unfold preserves_def stable_def constrains_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 162 | apply (subgoal_tac "s \<in> state & t \<in> state") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 163 | prefer 2 apply (blast dest!: Acts_type [THEN subsetD], force) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 164 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 165 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 166 | lemma Join_preserves [iff]: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 167 | "(F Join G \<in> preserves(v)) <-> | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 168 | (programify(F) \<in> preserves(v) & programify(G) \<in> preserves(v))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 169 | by (auto simp add: preserves_def INT_iff) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 170 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 171 | lemma JN_preserves [iff]: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 172 | "(JOIN(I,F): preserves(v)) <-> (\<forall>i \<in> I. programify(F(i)):preserves(v))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 173 | by (auto simp add: JN_stable preserves_def INT_iff) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 174 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 175 | lemma SKIP_preserves [iff]: "SKIP \<in> preserves(v)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 176 | by (auto simp add: preserves_def INT_iff) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 177 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 178 | lemma fun_pair_apply [simp]: "fun_pair(f,g,x) = <f(x), g(x)>" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 179 | apply (unfold fun_pair_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 180 | apply (simp (no_asm)) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 181 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 182 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 183 | lemma preserves_fun_pair: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 184 | "preserves(fun_pair(v,w)) = preserves(v) Int preserves(w)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 185 | apply (rule equalityI) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 186 | apply (auto simp add: preserves_def stable_def constrains_def, blast+) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 187 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 188 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 189 | lemma preserves_fun_pair_iff [iff]: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 190 | "F \<in> preserves(fun_pair(v, w)) <-> F \<in> preserves(v) Int preserves(w)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 191 | by (simp add: preserves_fun_pair) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 192 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 193 | lemma fun_pair_comp_distrib: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 194 | "(fun_pair(f, g) comp h)(x) = fun_pair(f comp h, g comp h, x)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 195 | by (simp add: fun_pair_def metacomp_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 196 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 197 | lemma comp_apply [simp]: "(f comp g)(x) = f(g(x))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 198 | by (simp add: metacomp_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 199 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 200 | lemma preserves_type: "preserves(v)<=program" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 201 | by (unfold preserves_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 202 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 203 | lemma preserves_into_program [TC]: "F \<in> preserves(f) ==> F \<in> program" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 204 | by (blast intro: preserves_type [THEN subsetD]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 205 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 206 | lemma subset_preserves_comp: "preserves(f) <= preserves(g comp f)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 207 | apply (auto simp add: preserves_def stable_def constrains_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 208 | apply (drule_tac x = "f (xb)" in spec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 209 | apply (drule_tac x = act in bspec, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 210 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 211 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 212 | lemma imp_preserves_comp: "F \<in> preserves(f) ==> F \<in> preserves(g comp f)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 213 | by (blast intro: subset_preserves_comp [THEN subsetD]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 214 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 215 | lemma preserves_subset_stable: "preserves(f) <= stable({s \<in> state. P(f(s))})"
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 216 | apply (auto simp add: preserves_def stable_def constrains_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 217 | apply (rename_tac s' s) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 218 | apply (subgoal_tac "f (s) = f (s') ") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 219 | apply (force+) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 220 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 221 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 222 | lemma preserves_imp_stable: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 223 |      "F \<in> preserves(f) ==> F \<in> stable({s \<in> state. P(f(s))})"
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 224 | by (blast intro: preserves_subset_stable [THEN subsetD]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 225 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 226 | lemma preserves_imp_increasing: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 227 | "[| F \<in> preserves(f); \<forall>x \<in> state. f(x):A |] ==> F \<in> Increasing.increasing(A, r, f)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 228 | apply (unfold Increasing.increasing_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 229 | apply (auto intro: preserves_into_program) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 230 | apply (rule_tac P = "%x. <k, x>:r" in preserves_imp_stable, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 231 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 232 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 233 | lemma preserves_id_subset_stable: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 234 | "st_set(A) ==> preserves(%x. x) <= stable(A)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 235 | apply (unfold preserves_def stable_def constrains_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 236 | apply (drule_tac x = xb in spec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 237 | apply (drule_tac x = act in bspec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 238 | apply (auto dest: ActsD) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 239 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 240 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 241 | lemma preserves_id_imp_stable: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 242 | "[| F \<in> preserves(%x. x); st_set(A) |] ==> F \<in> stable(A)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 243 | by (blast intro: preserves_id_subset_stable [THEN subsetD]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 244 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 245 | (** Added by Sidi **) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 246 | (** component_of **) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 247 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 248 | (* component_of is stronger than component *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 249 | lemma component_of_imp_component: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 250 | "F component_of H ==> F component H" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 251 | apply (unfold component_def component_of_def, blast) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 252 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 253 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 254 | (* component_of satisfies many of component's properties *) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 255 | lemma component_of_refl [simp]: "F \<in> program ==> F component_of F" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 256 | apply (unfold component_of_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 257 | apply (rule_tac x = SKIP in exI, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 258 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 259 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 260 | lemma component_of_SKIP [simp]: "F \<in> program ==>SKIP component_of F" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 261 | apply (unfold component_of_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 262 | apply (rule_tac x = F in exI, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 263 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 264 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 265 | lemma component_of_trans: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 266 | "[| F component_of G; G component_of H |] ==> F component_of H" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 267 | apply (unfold component_of_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 268 | apply (blast intro: Join_assoc [symmetric]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 269 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 270 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 271 | (** localize **) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 272 | lemma localize_Init_eq [simp]: "Init(localize(v,F)) = Init(F)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 273 | by (unfold localize_def, simp) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 274 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 275 | lemma localize_Acts_eq [simp]: "Acts(localize(v,F)) = Acts(F)" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 276 | by (unfold localize_def, simp) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 277 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 278 | lemma localize_AllowedActs_eq [simp]: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 279 | "AllowedActs(localize(v,F)) = AllowedActs(F) Int (\<Union>G \<in> preserves(v). Acts(G))" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 280 | apply (unfold localize_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 281 | apply (rule equalityI) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 282 | apply (auto dest: Acts_type [THEN subsetD]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 283 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 284 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 285 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 286 | (** Theorems used in ClientImpl **) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 287 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 288 | lemma stable_localTo_stable2: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 289 |  "[| F \<in> stable({s \<in> state. P(f(s), g(s))});  G \<in> preserves(f);  G \<in> preserves(g) |]  
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 290 |       ==> F Join G \<in> stable({s \<in> state. P(f(s), g(s))})"
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 291 | apply (auto dest: ActsD preserves_into_program simp add: stable_def constrains_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 292 | apply (case_tac "act \<in> Acts (F) ") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 293 | apply auto | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 294 | apply (drule preserves_imp_eq) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 295 | apply (drule_tac [3] preserves_imp_eq, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 296 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 297 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 298 | lemma Increasing_preserves_Stable: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 299 |      "[| F \<in> stable({s \<in> state. <f(s), g(s)>:r});  G \<in> preserves(f);    
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 300 | F Join G \<in> Increasing(A, r, g); | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 301 | \<forall>x \<in> state. f(x):A & g(x):A |] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 302 |       ==> F Join G \<in> Stable({s \<in> state. <f(s), g(s)>:r})"
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 303 | apply (auto simp add: stable_def Stable_def Increasing_def Constrains_def all_conj_distrib) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 304 | apply (simp_all add: constrains_type [THEN subsetD] preserves_type [THEN subsetD]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 305 | apply (blast intro: constrains_weaken) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 306 | (*The G case remains*) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 307 | apply (auto dest: ActsD simp add: preserves_def stable_def constrains_def ball_conj_distrib all_conj_distrib) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 308 | (*We have a G-action, so delete assumptions about F-actions*) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 309 | apply (erule_tac V = "\<forall>act \<in> Acts (F) . ?P (act)" in thin_rl) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 310 | apply (erule_tac V = "\<forall>k\<in>A. \<forall>act \<in> Acts (F) . ?P (k,act)" in thin_rl) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 311 | apply (subgoal_tac "f (x) = f (xa) ") | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 312 | apply (auto dest!: bspec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 313 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 314 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 315 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 316 | (** Lemma used in AllocImpl **) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 317 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 318 | lemma Constrains_UN_left: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 319 | "[| \<forall>x \<in> I. F \<in> A(x) Co B; F \<in> program |] ==> F:(\<Union>x \<in> I. A(x)) Co B" | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 320 | by (unfold Constrains_def constrains_def, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 321 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 322 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 323 | lemma stable_Join_Stable: | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 324 |  "[| F \<in> stable({s \<in> state. P(f(s), g(s))});  
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 325 |      \<forall>k \<in> A. F Join G \<in> Stable({s \<in> state. P(k, g(s))});  
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 326 | G \<in> preserves(f); \<forall>s \<in> state. f(s):A|] | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 327 |   ==> F Join G \<in> Stable({s \<in> state. P(f(s), g(s))})"
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 328 | apply (unfold stable_def Stable_def preserves_def) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 329 | apply (rule_tac A = "(\<Union>k \<in> A. {s \<in> state. f(s)=k} Int {s \<in> state. P (f (s), g (s))})" in Constrains_weaken_L)
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 330 | prefer 2 apply blast | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 331 | apply (rule Constrains_UN_left, auto) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 332 | apply (rule_tac A = "{s \<in> state. f(s)=k} Int {s \<in> state. P (f (s), g (s))} Int {s \<in> state. P (k, g (s))}" and A' = "({s \<in> state. f(s)=k} Un {s \<in> state. P (f (s), g (s))}) Int {s \<in> state. P (k, g (s))}" in Constrains_weaken)
 | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 333 | prefer 2 apply blast | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 334 | prefer 2 apply blast | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 335 | apply (rule Constrains_Int) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 336 | apply (rule constrains_imp_Constrains) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 337 | apply (auto simp add: constrains_type [THEN subsetD]) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 338 | apply (blast intro: constrains_weaken) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 339 | apply (drule_tac x = k in spec) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 340 | apply (blast intro: constrains_weaken) | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 341 | done | 
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 342 | |
| 
24382760fd89
converting more theories to Isar scripts, and tidying
 paulson parents: 
14046diff
changeset | 343 | end |