author | aspinall |
Fri, 30 Sep 2005 18:18:34 +0200 | |
changeset 17740 | fc385ce6187d |
parent 17256 | 526ff7cfd6ea |
child 17955 | 3b34516662c6 |
permissions | -rw-r--r-- |
3071 | 1 |
(* Title: HOLCF/IOA/meta_theory/ShortExecutions.thy |
3275 | 2 |
ID: $Id$ |
17256 | 3 |
Author: Olaf Müller |
4 |
*) |
|
3275 | 5 |
|
6 |
||
3361 | 7 |
fun thin_tac' j = |
8 |
rotate_tac (j - 1) THEN' |
|
17256 | 9 |
etac thin_rl THEN' |
3361 | 10 |
rotate_tac (~ (j - 1)); |
11 |
||
3275 | 12 |
|
13 |
||
14 |
(* ---------------------------------------------------------------- *) |
|
15 |
section "oraclebuild rewrite rules"; |
|
16 |
(* ---------------------------------------------------------------- *) |
|
17 |
||
18 |
||
17256 | 19 |
bind_thm ("oraclebuild_unfold", fix_prover2 (the_context ()) oraclebuild_def |
3275 | 20 |
"oraclebuild P = (LAM s t. case t of \ |
21 |
\ nil => nil\ |
|
22 |
\ | x##xs => \ |
|
17256 | 23 |
\ (case x of\ |
12028 | 24 |
\ UU => UU\ |
10835 | 25 |
\ | Def y => (Takewhile (%a.~ P a)$s)\ |
26 |
\ @@ (y>>(oraclebuild P$(TL$(Dropwhile (%a.~ P a)$s))$xs))\ |
|
3275 | 27 |
\ )\ |
28 |
\ )"); |
|
29 |
||
10835 | 30 |
Goal "oraclebuild P$sch$UU = UU"; |
3275 | 31 |
by (stac oraclebuild_unfold 1); |
32 |
by (Simp_tac 1); |
|
33 |
qed"oraclebuild_UU"; |
|
34 |
||
10835 | 35 |
Goal "oraclebuild P$sch$nil = nil"; |
3275 | 36 |
by (stac oraclebuild_unfold 1); |
37 |
by (Simp_tac 1); |
|
38 |
qed"oraclebuild_nil"; |
|
39 |
||
10835 | 40 |
Goal "oraclebuild P$s$(x>>t) = \ |
41 |
\ (Takewhile (%a.~ P a)$s) \ |
|
17256 | 42 |
\ @@ (x>>(oraclebuild P$(TL$(Dropwhile (%a.~ P a)$s))$t))"; |
3457 | 43 |
by (rtac trans 1); |
3275 | 44 |
by (stac oraclebuild_unfold 1); |
7229
6773ba0c36d5
renamed Cons to Consq in order to avoid clash with List.Cons;
wenzelm
parents:
6161
diff
changeset
|
45 |
by (asm_full_simp_tac (simpset() addsimps [Consq_def]) 1); |
6773ba0c36d5
renamed Cons to Consq in order to avoid clash with List.Cons;
wenzelm
parents:
6161
diff
changeset
|
46 |
by (simp_tac (simpset() addsimps [Consq_def]) 1); |
3275 | 47 |
qed"oraclebuild_cons"; |
48 |
||
49 |
||
50 |
||
17256 | 51 |
|
3275 | 52 |
(* ---------------------------------------------------------------- *) |
53 |
section "Cut rewrite rules"; |
|
54 |
(* ---------------------------------------------------------------- *) |
|
55 |
||
5068 | 56 |
Goalw [Cut_def] |
6161 | 57 |
"[| Forall (%a.~ P a) s; Finite s|] \ |
3275 | 58 |
\ ==> Cut P s =nil"; |
10835 | 59 |
by (subgoal_tac "Filter P$s = nil" 1); |
4098 | 60 |
by (asm_simp_tac (simpset() addsimps [oraclebuild_nil]) 1); |
3457 | 61 |
by (rtac ForallQFilterPnil 1); |
3275 | 62 |
by (REPEAT (atac 1)); |
63 |
qed"Cut_nil"; |
|
64 |
||
5068 | 65 |
Goalw [Cut_def] |
6161 | 66 |
"[| Forall (%a.~ P a) s; ~Finite s|] \ |
3275 | 67 |
\ ==> Cut P s =UU"; |
10835 | 68 |
by (subgoal_tac "Filter P$s= UU" 1); |
4098 | 69 |
by (asm_simp_tac (simpset() addsimps [oraclebuild_UU]) 1); |
3457 | 70 |
by (rtac ForallQFilterPUU 1); |
3275 | 71 |
by (REPEAT (atac 1)); |
72 |
qed"Cut_UU"; |
|
73 |
||
5068 | 74 |
Goalw [Cut_def] |
6161 | 75 |
"[| P t; Forall (%x.~ P x) ss; Finite ss|] \ |
3275 | 76 |
\ ==> Cut P (ss @@ (t>> rs)) \ |
77 |
\ = ss @@ (t >> Cut P rs)"; |
|
78 |
||
4098 | 79 |
by (asm_full_simp_tac (simpset() addsimps [ForallQFilterPnil,oraclebuild_cons, |
3275 | 80 |
TakewhileConc,DropwhileConc]) 1); |
81 |
qed"Cut_Cons"; |
|
82 |
||
83 |
||
84 |
(* ---------------------------------------------------------------- *) |
|
85 |
section "Cut lemmas for main theorem"; |
|
86 |
(* ---------------------------------------------------------------- *) |
|
87 |
||
88 |
||
10835 | 89 |
Goal "Filter P$s = Filter P$(Cut P s)"; |
3275 | 90 |
|
3847 | 91 |
by (res_inst_tac [("A1","%x. True") |
3275 | 92 |
,("Q1","%x.~ P x"), ("x1","s")] |
93 |
(take_lemma_induct RS mp) 1); |
|
94 |
by (Fast_tac 3); |
|
95 |
by (case_tac "Finite s" 1); |
|
4098 | 96 |
by (asm_full_simp_tac (simpset() addsimps [Cut_nil, |
3275 | 97 |
ForallQFilterPnil]) 1); |
4098 | 98 |
by (asm_full_simp_tac (simpset() addsimps [Cut_UU, |
3275 | 99 |
ForallQFilterPUU]) 1); |
100 |
(* main case *) |
|
4098 | 101 |
by (asm_full_simp_tac (simpset() addsimps [Cut_Cons,ForallQFilterPnil]) 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4283
diff
changeset
|
102 |
by Auto_tac; |
3275 | 103 |
qed"FilterCut"; |
104 |
||
105 |
||
5068 | 106 |
Goal "Cut P (Cut P s) = (Cut P s)"; |
3275 | 107 |
|
3847 | 108 |
by (res_inst_tac [("A1","%x. True") |
3275 | 109 |
,("Q1","%x.~ P x"), ("x1","s")] |
110 |
(take_lemma_less_induct RS mp) 1); |
|
111 |
by (Fast_tac 3); |
|
112 |
by (case_tac "Finite s" 1); |
|
4098 | 113 |
by (asm_full_simp_tac (simpset() addsimps [Cut_nil, |
3275 | 114 |
ForallQFilterPnil]) 1); |
4098 | 115 |
by (asm_full_simp_tac (simpset() addsimps [Cut_UU, |
3275 | 116 |
ForallQFilterPUU]) 1); |
117 |
(* main case *) |
|
4098 | 118 |
by (asm_full_simp_tac (simpset() addsimps [Cut_Cons,ForallQFilterPnil]) 1); |
3457 | 119 |
by (rtac take_reduction 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4283
diff
changeset
|
120 |
by Auto_tac; |
3275 | 121 |
qed"Cut_idemp"; |
122 |
||
123 |
||
10835 | 124 |
Goal "Map f$(Cut (P o f) s) = Cut P (Map f$s)"; |
3275 | 125 |
|
3847 | 126 |
by (res_inst_tac [("A1","%x. True") |
3275 | 127 |
,("Q1","%x.~ P (f x)"), ("x1","s")] |
128 |
(take_lemma_less_induct RS mp) 1); |
|
129 |
by (Fast_tac 3); |
|
130 |
by (case_tac "Finite s" 1); |
|
17256 | 131 |
by (asm_full_simp_tac (simpset() addsimps [Cut_nil]) 1); |
3457 | 132 |
by (rtac (Cut_nil RS sym) 1); |
17256 | 133 |
by (asm_full_simp_tac (simpset() addsimps [ForallMap,o_def]) 1); |
4098 | 134 |
by (asm_full_simp_tac (simpset() addsimps [Map2Finite]) 1); |
3275 | 135 |
(* csae ~ Finite s *) |
4098 | 136 |
by (asm_full_simp_tac (simpset() addsimps [Cut_UU]) 1); |
16744 | 137 |
by (rtac (Cut_UU (*RS sym*)) 1); |
17256 | 138 |
by (asm_full_simp_tac (simpset() addsimps [ForallMap,o_def]) 1); |
4098 | 139 |
by (asm_full_simp_tac (simpset() addsimps [Map2Finite]) 1); |
3275 | 140 |
(* main case *) |
4098 | 141 |
by (asm_full_simp_tac (simpset() addsimps [Cut_Cons,MapConc, |
3275 | 142 |
ForallMap,FiniteMap1,o_def]) 1); |
3457 | 143 |
by (rtac take_reduction 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4283
diff
changeset
|
144 |
by Auto_tac; |
3275 | 145 |
qed"MapCut"; |
146 |
||
147 |
||
5068 | 148 |
Goal "~Finite s --> Cut P s << s"; |
3361 | 149 |
by (strip_tac 1); |
3457 | 150 |
by (rtac (take_lemma_less RS iffD1) 1); |
3361 | 151 |
by (strip_tac 1); |
3457 | 152 |
by (rtac mp 1); |
153 |
by (assume_tac 2); |
|
3361 | 154 |
by (thin_tac' 1 1); |
155 |
by (res_inst_tac [("x","s")] spec 1); |
|
9877 | 156 |
by (rtac nat_less_induct 1); |
3361 | 157 |
by (strip_tac 1); |
158 |
ren "na n s" 1; |
|
159 |
by (case_tac "Forall (%x. ~ P x) s" 1); |
|
3457 | 160 |
by (rtac (take_lemma_less RS iffD2 RS spec) 1); |
4098 | 161 |
by (asm_full_simp_tac (simpset() addsimps [Cut_UU]) 1); |
3361 | 162 |
(* main case *) |
3457 | 163 |
by (dtac divide_Seq3 1); |
3361 | 164 |
by (REPEAT (etac exE 1)); |
165 |
by (REPEAT (etac conjE 1)); |
|
166 |
by (hyp_subst_tac 1); |
|
4098 | 167 |
by (asm_full_simp_tac (simpset() addsimps [Cut_Cons]) 1); |
3457 | 168 |
by (rtac take_reduction_less 1); |
3361 | 169 |
(* auto makes also reasoning about Finiteness of parts of s ! *) |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4283
diff
changeset
|
170 |
by Auto_tac; |
3361 | 171 |
qed_spec_mp"Cut_prefixcl_nFinite"; |
172 |
||
173 |
||
174 |
||
5068 | 175 |
Goal "!!ex .is_exec_frag A (s,ex) ==> is_exec_frag A (s,Cut P ex)"; |
3361 | 176 |
by (case_tac "Finite ex" 1); |
177 |
by (cut_inst_tac [("s","ex"),("P","P")] Cut_prefixcl_Finite 1); |
|
3457 | 178 |
by (assume_tac 1); |
179 |
by (etac exE 1); |
|
180 |
by (rtac exec_prefix2closed 1); |
|
3361 | 181 |
by (eres_inst_tac [("s","ex"),("t","Cut P ex @@ y")] subst 1); |
3457 | 182 |
by (assume_tac 1); |
183 |
by (etac exec_prefixclosed 1); |
|
184 |
by (etac Cut_prefixcl_nFinite 1); |
|
3361 | 185 |
qed"execThruCut"; |
186 |
||
3275 | 187 |
|
188 |
||
189 |
(* ---------------------------------------------------------------- *) |
|
190 |
section "Main Cut Theorem"; |
|
191 |
(* ---------------------------------------------------------------- *) |
|
192 |
||
193 |
||
194 |
||
5068 | 195 |
Goalw [schedules_def,has_schedule_def] |
10835 | 196 |
"[|sch : schedules A ; tr = Filter (%a. a:ext A)$sch|] \ |
3275 | 197 |
\ ==> ? sch. sch : schedules A & \ |
10835 | 198 |
\ tr = Filter (%a. a:ext A)$sch &\ |
3275 | 199 |
\ LastActExtsch A sch"; |
200 |
||
201 |
by (safe_tac set_cs); |
|
10835 | 202 |
by (res_inst_tac [("x","filter_act$(Cut (%a. fst a:ext A) (snd ex))")] exI 1); |
4098 | 203 |
by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1); |
3275 | 204 |
by (pair_tac "ex" 1); |
205 |
by (safe_tac set_cs); |
|
3847 | 206 |
by (res_inst_tac [("x","(x,Cut (%a. fst a:ext A) y)")] exI 1); |
3275 | 207 |
by (Asm_simp_tac 1); |
208 |
||
209 |
(* Subgoal 1: Lemma: propagation of execution through Cut *) |
|
210 |
||
4098 | 211 |
by (asm_full_simp_tac (simpset() addsimps [execThruCut]) 1); |
3275 | 212 |
|
213 |
(* Subgoal 2: Lemma: Filter P s = Filter P (Cut P s) *) |
|
214 |
||
4098 | 215 |
by (simp_tac (simpset() addsimps [filter_act_def]) 1); |
10835 | 216 |
by (subgoal_tac "Map fst$(Cut (%a. fst a: ext A) y)= Cut (%a. a:ext A) (Map fst$y)" 1); |
4283 | 217 |
|
3457 | 218 |
by (rtac (rewrite_rule [o_def] MapCut) 2); |
4098 | 219 |
by (asm_full_simp_tac (simpset() addsimps [FilterCut RS sym]) 1); |
3275 | 220 |
|
221 |
(* Subgoal 3: Lemma: Cut idempotent *) |
|
222 |
||
4098 | 223 |
by (simp_tac (simpset() addsimps [LastActExtsch_def,filter_act_def]) 1); |
10835 | 224 |
by (subgoal_tac "Map fst$(Cut (%a. fst a: ext A) y)= Cut (%a. a:ext A) (Map fst$y)" 1); |
3457 | 225 |
by (rtac (rewrite_rule [o_def] MapCut) 2); |
4098 | 226 |
by (asm_full_simp_tac (simpset() addsimps [Cut_idemp]) 1); |
3275 | 227 |
qed"exists_LastActExtsch"; |
228 |
||
229 |
||
230 |
(* ---------------------------------------------------------------- *) |
|
231 |
section "Further Cut lemmas"; |
|
232 |
(* ---------------------------------------------------------------- *) |
|
233 |
||
5068 | 234 |
Goalw [LastActExtsch_def] |
10835 | 235 |
"[| LastActExtsch A sch; Filter (%x. x:ext A)$sch = nil |] \ |
3275 | 236 |
\ ==> sch=nil"; |
3457 | 237 |
by (dtac FilternPnilForallP 1); |
238 |
by (etac conjE 1); |
|
239 |
by (dtac Cut_nil 1); |
|
240 |
by (assume_tac 1); |
|
3275 | 241 |
by (Asm_full_simp_tac 1); |
242 |
qed"LastActExtimplnil"; |
|
243 |
||
5068 | 244 |
Goalw [LastActExtsch_def] |
10835 | 245 |
"[| LastActExtsch A sch; Filter (%x. x:ext A)$sch = UU |] \ |
3275 | 246 |
\ ==> sch=UU"; |
3457 | 247 |
by (dtac FilternPUUForallP 1); |
248 |
by (etac conjE 1); |
|
249 |
by (dtac Cut_UU 1); |
|
250 |
by (assume_tac 1); |
|
3275 | 251 |
by (Asm_full_simp_tac 1); |
252 |
qed"LastActExtimplUU"; |