| 3071 |      1 | (*  Title:      HOLCF/IOA/meta_theory/ShortExecutions.thy
 | 
| 3275 |      2 |     ID:         $Id$
 | 
| 12218 |      3 |     Author:     Olaf Müller
 | 
| 17256 |      4 | *)
 | 
| 3071 |      5 | 
 | 
| 17256 |      6 | theory ShortExecutions
 | 
|  |      7 | imports Traces
 | 
|  |      8 | begin
 | 
| 3071 |      9 | 
 | 
| 17256 |     10 | text {*
 | 
|  |     11 |   Some properties about @{text "Cut ex"}, defined as follows:
 | 
| 3071 |     12 | 
 | 
| 17256 |     13 |   For every execution ex there is another shorter execution @{text "Cut ex"}
 | 
|  |     14 |   that has the same trace as ex, but its schedule ends with an external action.
 | 
|  |     15 | *}
 | 
| 3071 |     16 | 
 | 
| 17256 |     17 | consts
 | 
| 3071 |     18 | 
 | 
| 4283 |     19 | (*  LastActExtex      ::"('a,'s)ioa => ('a,'s) pairs  => bool"*)
 | 
| 3275 |     20 |   LastActExtsch     ::"('a,'s)ioa => 'a Seq         => bool"
 | 
| 3071 |     21 | 
 | 
| 3275 |     22 |   Cut               ::"('a => bool) => 'a Seq    => 'a Seq"
 | 
| 17256 |     23 | 
 | 
| 3275 |     24 |   oraclebuild       ::"('a => bool) => 'a Seq -> 'a Seq -> 'a Seq"
 | 
| 3071 |     25 | 
 | 
|  |     26 | defs
 | 
|  |     27 | 
 | 
| 17256 |     28 | LastActExtsch_def:
 | 
| 3847 |     29 |   "LastActExtsch A sch == (Cut (%x. x: ext A) sch = sch)"
 | 
| 3275 |     30 | 
 | 
| 17256 |     31 | (* LastActExtex_def:
 | 
| 10835 |     32 |   "LastActExtex A ex == LastActExtsch A (filter_act$ex)" *)
 | 
| 3275 |     33 | 
 | 
| 17256 |     34 | Cut_def:
 | 
| 10835 |     35 |   "Cut P s == oraclebuild P$s$(Filter P$s)"
 | 
| 3275 |     36 | 
 | 
| 17256 |     37 | oraclebuild_def:
 | 
|  |     38 |   "oraclebuild P == (fix$(LAM h s t. case t of
 | 
| 3275 |     39 |        nil => nil
 | 
| 17256 |     40 |     | x##xs =>
 | 
|  |     41 |       (case x of
 | 
| 12028 |     42 |         UU => UU
 | 
| 10835 |     43 |       | Def y => (Takewhile (%x.~P x)$s)
 | 
|  |     44 |                   @@ (y>>(h$(TL$(Dropwhile (%x.~ P x)$s))$xs))
 | 
| 3275 |     45 |       )
 | 
|  |     46 |     ))"
 | 
|  |     47 | 
 | 
|  |     48 | 
 | 
| 17256 |     49 | axioms
 | 
| 3071 |     50 | 
 | 
| 17256 |     51 | Cut_prefixcl_Finite:
 | 
| 3361 |     52 |   "Finite s ==> (? y. s = Cut P s @@ y)"
 | 
| 3071 |     53 | 
 | 
| 17256 |     54 | LastActExtsmall1:
 | 
| 10835 |     55 |   "LastActExtsch A sch ==> LastActExtsch A (TL$(Dropwhile P$sch))"
 | 
| 3071 |     56 | 
 | 
| 17256 |     57 | LastActExtsmall2:
 | 
| 3275 |     58 |   "[| Finite sch1; LastActExtsch A (sch1 @@ sch2) |] ==> LastActExtsch A sch2"
 | 
| 3071 |     59 | 
 | 
| 19741 |     60 | 
 | 
|  |     61 | 
 | 
|  |     62 | ML {*
 | 
|  |     63 | fun thin_tac' j =
 | 
|  |     64 |   rotate_tac (j - 1) THEN'
 | 
|  |     65 |   etac thin_rl THEN'
 | 
|  |     66 |   rotate_tac (~ (j - 1))
 | 
|  |     67 | *}
 | 
|  |     68 | 
 | 
|  |     69 | 
 | 
|  |     70 | subsection "oraclebuild rewrite rules"
 | 
|  |     71 | 
 | 
|  |     72 | 
 | 
|  |     73 | lemma oraclebuild_unfold:
 | 
|  |     74 | "oraclebuild P = (LAM s t. case t of 
 | 
|  |     75 |        nil => nil
 | 
|  |     76 |     | x##xs => 
 | 
|  |     77 |       (case x of
 | 
|  |     78 |         UU => UU
 | 
|  |     79 |       | Def y => (Takewhile (%a.~ P a)$s)
 | 
|  |     80 |                   @@ (y>>(oraclebuild P$(TL$(Dropwhile (%a.~ P a)$s))$xs))
 | 
|  |     81 |       )
 | 
|  |     82 |     )"
 | 
|  |     83 | apply (rule trans)
 | 
|  |     84 | apply (rule fix_eq2)
 | 
|  |     85 | apply (rule oraclebuild_def)
 | 
|  |     86 | apply (rule beta_cfun)
 | 
|  |     87 | apply simp
 | 
|  |     88 | done
 | 
|  |     89 | 
 | 
|  |     90 | lemma oraclebuild_UU: "oraclebuild P$sch$UU = UU"
 | 
|  |     91 | apply (subst oraclebuild_unfold)
 | 
|  |     92 | apply simp
 | 
|  |     93 | done
 | 
|  |     94 | 
 | 
|  |     95 | lemma oraclebuild_nil: "oraclebuild P$sch$nil = nil"
 | 
|  |     96 | apply (subst oraclebuild_unfold)
 | 
|  |     97 | apply simp
 | 
|  |     98 | done
 | 
|  |     99 | 
 | 
|  |    100 | lemma oraclebuild_cons: "oraclebuild P$s$(x>>t) =  
 | 
|  |    101 |           (Takewhile (%a.~ P a)$s)    
 | 
|  |    102 |            @@ (x>>(oraclebuild P$(TL$(Dropwhile (%a.~ P a)$s))$t))"
 | 
|  |    103 | apply (rule trans)
 | 
|  |    104 | apply (subst oraclebuild_unfold)
 | 
|  |    105 | apply (simp add: Consq_def)
 | 
|  |    106 | apply (simp add: Consq_def)
 | 
|  |    107 | done
 | 
|  |    108 | 
 | 
|  |    109 | 
 | 
|  |    110 | subsection "Cut rewrite rules"
 | 
|  |    111 | 
 | 
|  |    112 | lemma Cut_nil: 
 | 
|  |    113 | "[| Forall (%a.~ P a) s; Finite s|]  
 | 
|  |    114 |             ==> Cut P s =nil"
 | 
|  |    115 | apply (unfold Cut_def)
 | 
|  |    116 | apply (subgoal_tac "Filter P$s = nil")
 | 
|  |    117 | apply (simp (no_asm_simp) add: oraclebuild_nil)
 | 
|  |    118 | apply (rule ForallQFilterPnil)
 | 
|  |    119 | apply assumption+
 | 
|  |    120 | done
 | 
|  |    121 | 
 | 
|  |    122 | lemma Cut_UU: 
 | 
|  |    123 | "[| Forall (%a.~ P a) s; ~Finite s|]  
 | 
|  |    124 |             ==> Cut P s =UU"
 | 
|  |    125 | apply (unfold Cut_def)
 | 
|  |    126 | apply (subgoal_tac "Filter P$s= UU")
 | 
|  |    127 | apply (simp (no_asm_simp) add: oraclebuild_UU)
 | 
|  |    128 | apply (rule ForallQFilterPUU)
 | 
|  |    129 | apply assumption+
 | 
|  |    130 | done
 | 
|  |    131 | 
 | 
|  |    132 | lemma Cut_Cons: 
 | 
|  |    133 | "[| P t;  Forall (%x.~ P x) ss; Finite ss|]  
 | 
|  |    134 |             ==> Cut P (ss @@ (t>> rs))  
 | 
|  |    135 |                  = ss @@ (t >> Cut P rs)"
 | 
|  |    136 | apply (unfold Cut_def)
 | 
|  |    137 | apply (simp add: ForallQFilterPnil oraclebuild_cons TakewhileConc DropwhileConc)
 | 
|  |    138 | done
 | 
|  |    139 | 
 | 
|  |    140 | 
 | 
|  |    141 | subsection "Cut lemmas for main theorem"
 | 
|  |    142 | 
 | 
|  |    143 | lemma FilterCut: "Filter P$s = Filter P$(Cut P s)"
 | 
|  |    144 | apply (rule_tac A1 = "%x. True" and Q1 = "%x.~ P x" and x1 = "s" in take_lemma_induct [THEN mp])
 | 
|  |    145 | prefer 3 apply (fast)
 | 
|  |    146 | apply (case_tac "Finite s")
 | 
|  |    147 | apply (simp add: Cut_nil ForallQFilterPnil)
 | 
|  |    148 | apply (simp add: Cut_UU ForallQFilterPUU)
 | 
|  |    149 | (* main case *)
 | 
|  |    150 | apply (simp add: Cut_Cons ForallQFilterPnil)
 | 
|  |    151 | done
 | 
|  |    152 | 
 | 
|  |    153 | 
 | 
|  |    154 | lemma Cut_idemp: "Cut P (Cut P s) = (Cut P s)"
 | 
|  |    155 | apply (rule_tac A1 = "%x. True" and Q1 = "%x.~ P x" and x1 = "s" in
 | 
|  |    156 |   take_lemma_less_induct [THEN mp])
 | 
|  |    157 | prefer 3 apply (fast)
 | 
|  |    158 | apply (case_tac "Finite s")
 | 
|  |    159 | apply (simp add: Cut_nil ForallQFilterPnil)
 | 
|  |    160 | apply (simp add: Cut_UU ForallQFilterPUU)
 | 
|  |    161 | (* main case *)
 | 
|  |    162 | apply (simp add: Cut_Cons ForallQFilterPnil)
 | 
|  |    163 | apply (rule take_reduction)
 | 
|  |    164 | apply auto
 | 
|  |    165 | done
 | 
|  |    166 | 
 | 
|  |    167 | 
 | 
|  |    168 | lemma MapCut: "Map f$(Cut (P o f) s) = Cut P (Map f$s)"
 | 
|  |    169 | apply (rule_tac A1 = "%x. True" and Q1 = "%x.~ P (f x) " and x1 = "s" in
 | 
|  |    170 |   take_lemma_less_induct [THEN mp])
 | 
|  |    171 | prefer 3 apply (fast)
 | 
|  |    172 | apply (case_tac "Finite s")
 | 
|  |    173 | apply (simp add: Cut_nil)
 | 
|  |    174 | apply (rule Cut_nil [symmetric])
 | 
|  |    175 | apply (simp add: ForallMap o_def)
 | 
|  |    176 | apply (simp add: Map2Finite)
 | 
|  |    177 | (* csae ~ Finite s *)
 | 
|  |    178 | apply (simp add: Cut_UU)
 | 
|  |    179 | apply (rule Cut_UU)
 | 
|  |    180 | apply (simp add: ForallMap o_def)
 | 
|  |    181 | apply (simp add: Map2Finite)
 | 
|  |    182 | (* main case *)
 | 
|  |    183 | apply (simp add: Cut_Cons MapConc ForallMap FiniteMap1 o_def)
 | 
|  |    184 | apply (rule take_reduction)
 | 
|  |    185 | apply auto
 | 
|  |    186 | done
 | 
|  |    187 | 
 | 
|  |    188 | 
 | 
|  |    189 | lemma Cut_prefixcl_nFinite [rule_format (no_asm)]: "~Finite s --> Cut P s << s"
 | 
|  |    190 | apply (intro strip)
 | 
|  |    191 | apply (rule take_lemma_less [THEN iffD1])
 | 
|  |    192 | apply (intro strip)
 | 
|  |    193 | apply (rule mp)
 | 
|  |    194 | prefer 2 apply (assumption)
 | 
|  |    195 | apply (tactic "thin_tac' 1 1")
 | 
|  |    196 | apply (rule_tac x = "s" in spec)
 | 
|  |    197 | apply (rule nat_less_induct)
 | 
|  |    198 | apply (intro strip)
 | 
|  |    199 | apply (rename_tac na n s)
 | 
|  |    200 | apply (case_tac "Forall (%x. ~ P x) s")
 | 
|  |    201 | apply (rule take_lemma_less [THEN iffD2, THEN spec])
 | 
|  |    202 | apply (simp add: Cut_UU)
 | 
|  |    203 | (* main case *)
 | 
|  |    204 | apply (drule divide_Seq3)
 | 
|  |    205 | apply (erule exE)+
 | 
|  |    206 | apply (erule conjE)+
 | 
|  |    207 | apply hypsubst
 | 
|  |    208 | apply (simp add: Cut_Cons)
 | 
|  |    209 | apply (rule take_reduction_less)
 | 
|  |    210 | (* auto makes also reasoning about Finiteness of parts of s ! *)
 | 
|  |    211 | apply auto
 | 
|  |    212 | done
 | 
|  |    213 | 
 | 
|  |    214 | 
 | 
|  |    215 | lemma execThruCut: "!!ex .is_exec_frag A (s,ex) ==> is_exec_frag A (s,Cut P ex)"
 | 
|  |    216 | apply (case_tac "Finite ex")
 | 
|  |    217 | apply (cut_tac s = "ex" and P = "P" in Cut_prefixcl_Finite)
 | 
|  |    218 | apply assumption
 | 
|  |    219 | apply (erule exE)
 | 
|  |    220 | apply (rule exec_prefix2closed)
 | 
|  |    221 | apply (erule_tac s = "ex" and t = "Cut P ex @@ y" in subst)
 | 
|  |    222 | apply assumption
 | 
|  |    223 | apply (erule exec_prefixclosed)
 | 
|  |    224 | apply (erule Cut_prefixcl_nFinite)
 | 
|  |    225 | done
 | 
|  |    226 | 
 | 
|  |    227 | 
 | 
|  |    228 | subsection "Main Cut Theorem"
 | 
|  |    229 | 
 | 
|  |    230 | lemma exists_LastActExtsch: 
 | 
|  |    231 |  "[|sch : schedules A ; tr = Filter (%a. a:ext A)$sch|]  
 | 
|  |    232 |     ==> ? sch. sch : schedules A &  
 | 
|  |    233 |                tr = Filter (%a. a:ext A)$sch & 
 | 
|  |    234 |                LastActExtsch A sch"
 | 
|  |    235 | 
 | 
|  |    236 | apply (unfold schedules_def has_schedule_def)
 | 
|  |    237 | apply (tactic "safe_tac set_cs")
 | 
|  |    238 | apply (rule_tac x = "filter_act$ (Cut (%a. fst a:ext A) (snd ex))" in exI)
 | 
|  |    239 | apply (simp add: executions_def)
 | 
|  |    240 | apply (tactic {* pair_tac "ex" 1 *})
 | 
|  |    241 | apply (tactic "safe_tac set_cs")
 | 
|  |    242 | apply (rule_tac x = " (x,Cut (%a. fst a:ext A) y) " in exI)
 | 
|  |    243 | apply (simp (no_asm_simp))
 | 
|  |    244 | 
 | 
|  |    245 | (* Subgoal 1: Lemma:  propagation of execution through Cut *)
 | 
|  |    246 | 
 | 
|  |    247 | apply (simp add: execThruCut)
 | 
|  |    248 | 
 | 
|  |    249 | (* Subgoal 2:  Lemma:  Filter P s = Filter P (Cut P s) *)
 | 
|  |    250 | 
 | 
|  |    251 | apply (simp (no_asm) add: filter_act_def)
 | 
|  |    252 | apply (subgoal_tac "Map fst$ (Cut (%a. fst a: ext A) y) = Cut (%a. a:ext A) (Map fst$y) ")
 | 
|  |    253 | 
 | 
|  |    254 | apply (rule_tac [2] MapCut [unfolded o_def])
 | 
|  |    255 | apply (simp add: FilterCut [symmetric])
 | 
|  |    256 | 
 | 
|  |    257 | (* Subgoal 3: Lemma: Cut idempotent  *)
 | 
|  |    258 | 
 | 
|  |    259 | apply (simp (no_asm) add: LastActExtsch_def filter_act_def)
 | 
|  |    260 | apply (subgoal_tac "Map fst$ (Cut (%a. fst a: ext A) y) = Cut (%a. a:ext A) (Map fst$y) ")
 | 
|  |    261 | apply (rule_tac [2] MapCut [unfolded o_def])
 | 
|  |    262 | apply (simp add: Cut_idemp)
 | 
|  |    263 | done
 | 
|  |    264 | 
 | 
|  |    265 | 
 | 
|  |    266 | subsection "Further Cut lemmas"
 | 
|  |    267 | 
 | 
|  |    268 | lemma LastActExtimplnil: 
 | 
|  |    269 |   "[| LastActExtsch A sch; Filter (%x. x:ext A)$sch = nil |]  
 | 
|  |    270 |     ==> sch=nil"
 | 
|  |    271 | apply (unfold LastActExtsch_def)
 | 
|  |    272 | apply (drule FilternPnilForallP)
 | 
|  |    273 | apply (erule conjE)
 | 
|  |    274 | apply (drule Cut_nil)
 | 
|  |    275 | apply assumption
 | 
|  |    276 | apply simp
 | 
|  |    277 | done
 | 
|  |    278 | 
 | 
|  |    279 | lemma LastActExtimplUU: 
 | 
|  |    280 |   "[| LastActExtsch A sch; Filter (%x. x:ext A)$sch = UU |]  
 | 
|  |    281 |     ==> sch=UU"
 | 
|  |    282 | apply (unfold LastActExtsch_def)
 | 
|  |    283 | apply (drule FilternPUUForallP)
 | 
|  |    284 | apply (erule conjE)
 | 
|  |    285 | apply (drule Cut_UU)
 | 
|  |    286 | apply assumption
 | 
|  |    287 | apply simp
 | 
|  |    288 | done
 | 
| 17256 |    289 | 
 | 
| 3071 |    290 | end
 |