src/HOL/MicroJava/BV/Kildall.thy
author kleing
Tue, 05 Dec 2000 14:08:56 +0100
changeset 10592 fc0b575a2cf7
parent 10496 f2d304bdf3cc
child 10661 fcd8d4e7d758
permissions -rw-r--r--
BCV Integration
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     1
(*  Title:      HOL/BCV/Kildall.thy
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     2
    ID:         $Id$
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     3
    Author:     Tobias Nipkow
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     4
    Copyright   2000 TUM
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     5
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     6
Kildall's algorithm
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     7
*)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     8
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     9
header "Kildall's Algorithm"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    10
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    11
theory Kildall = DFA_Framework:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    12
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    13
constdefs
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    14
 pres_type :: "(nat => 's => 's) => nat => 's set => bool"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    15
"pres_type step n A == !s:A. !p<n. step p s : A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    16
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    17
 mono :: "'s ord => (nat => 's => 's) => nat => 's set => bool"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    18
"mono r step n A ==
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    19
 !s p t. s : A & p < n & s <=_r t --> step p s <=_r step p t"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    20
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    21
consts
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    22
 iter :: "('s sl * (nat => 's => 's) * (nat => nat list))
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    23
          * 's list * nat set => 's list"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    24
 propa :: "'s binop => nat list => 's => 's list => nat set => 's list * nat set"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    25
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    26
primrec
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    27
"propa f []     t ss w = (ss,w)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    28
"propa f (q#qs) t ss w = (let u = t +_f ss!q;
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    29
                              w' = (if u = ss!q then w else insert q w)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    30
                          in propa f qs t (ss[q := u]) w')"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    31
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    32
recdef iter
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    33
 "same_fst (%((A,r,f),step,succs). semilat(A,r,f) & acc r)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    34
         (%((A,r,f),step,succs).
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    35
  {(ss',ss) . ss <[r] ss'} <*lex*> finite_psubset)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    36
"iter(((A,r,f),step,succs),ss,w) =
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    37
  (if semilat(A,r,f) & acc r & (!p:w. p < size ss) &
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    38
      bounded succs (size ss) & set ss <= A & pres_type step (size ss) A
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    39
   then if w={} then ss
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    40
        else let p = SOME p. p : w
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    41
             in iter(((A,r,f),step,succs),
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    42
                     propa f (succs p) (step p (ss!p)) ss (w-{p}))
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    43
   else arbitrary)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    44
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    45
constdefs
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    46
 unstables ::
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    47
 "'s binop => (nat => 's => 's) => (nat => nat list) => 's list => nat set"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    48
"unstables f step succs ss ==
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    49
 {p. p < size ss & (? q:set(succs p). step p (ss!p) +_f ss!q ~= ss!q)}"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    50
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    51
 kildall :: "'s sl => (nat => 's => 's) => (nat => nat list)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    52
             => 's list => 's list"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    53
"kildall Arf step succs ss ==
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    54
 iter((Arf,step,succs),ss,unstables (snd(snd Arf)) step succs ss)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    55
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    56
consts merges :: "'s binop => 's => nat list => 's list => 's list"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    57
primrec
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    58
"merges f s []     ss = ss"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    59
"merges f s (p#ps) ss = merges f s ps (ss[p := s +_f ss!p])"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    60
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    61
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    62
lemmas [simp] = Let_def le_iff_plus_unchanged [symmetric];
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    63
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    64
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    65
lemma pres_typeD:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    66
  "[| pres_type step n A; s:A; p<n |] ==> step p s : A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    67
  by (unfold pres_type_def, blast)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    68
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    69
lemma boundedD:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    70
  "[| bounded succs n; p < n; q : set(succs p) |] ==> q < n"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    71
  by (unfold bounded_def, blast)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    72
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    73
lemma monoD:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    74
  "[| mono r step n A; p < n; s:A; s <=_r t |] ==> step p s <=_r step p t"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    75
  by (unfold mono_def, blast)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    76
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    77
(** merges **)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    78
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    79
lemma length_merges [rule_format, simp]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    80
  "!ss. size(merges f t ps ss) = size ss"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    81
  by (induct_tac ps, auto)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    82
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    83
lemma merges_preserves_type [rule_format, simp]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    84
  "[| semilat(A,r,f) |] ==> 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    85
     !xs. xs : list n A --> x : A --> (!p : set ps. p<n) 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    86
          --> merges f x ps xs : list n A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    87
  apply (frule semilatDclosedI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    88
  apply (unfold closed_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    89
  apply (induct_tac ps)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    90
  apply auto
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    91
  done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    92
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    93
lemma merges_incr [rule_format]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    94
  "[| semilat(A,r,f) |] ==> 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    95
     !xs. xs : list n A --> x : A --> (!p:set ps. p<size xs) --> xs <=[r] merges f x ps xs"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    96
  apply (induct_tac ps)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    97
   apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    98
  apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    99
  apply clarify
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   100
  apply (rule order_trans)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   101
    apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   102
   apply (erule list_update_incr)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   103
     apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   104
    apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   105
   apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   106
  apply (blast intro!: listE_set intro: closedD listE_length [THEN nth_in])
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   107
  done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   108
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   109
lemma merges_same_conv [rule_format]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   110
  "[| semilat(A,r,f) |] ==> 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   111
     (!xs. xs : list n A --> x:A --> (!p:set ps. p<size xs) --> 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   112
     (merges f x ps xs = xs) = (!p:set ps. x <=_r xs!p))"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   113
  apply (induct_tac ps)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   114
   apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   115
  apply clarsimp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   116
  apply (rename_tac p ps xs)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   117
  apply (rule iffI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   118
   apply (rule context_conjI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   119
    apply (subgoal_tac "xs[p := x +_f xs!p] <=[r] xs")
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   120
     apply (force dest!: le_listD simp add: nth_list_update)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   121
    apply (erule subst, rule merges_incr)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   122
        apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   123
       apply (blast intro!: listE_set intro: closedD listE_length [THEN nth_in])
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   124
     apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   125
    apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   126
   apply clarify
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   127
   apply (rotate_tac -2)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   128
   apply (erule allE)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   129
   apply (erule impE)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   130
    apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   131
   apply (erule impE)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   132
    apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   133
   apply (drule bspec)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   134
    apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   135
   apply (simp add: le_iff_plus_unchanged [THEN iffD1] list_update_same_conv [THEN iffD2])
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   136
  apply clarify 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   137
  apply (simp add: le_iff_plus_unchanged [THEN iffD1] list_update_same_conv [THEN iffD2])
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   138
  done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   139
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   140
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   141
lemma list_update_le_listI [rule_format]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   142
  "set xs <= A --> set ys <= A --> xs <=[r] ys --> p < size xs -->  
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   143
   x <=_r ys!p --> semilat(A,r,f) --> x:A --> 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   144
   xs[p := x +_f xs!p] <=[r] ys"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   145
  apply (unfold Listn.le_def lesub_def semilat_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   146
  apply (simp add: list_all2_conv_all_nth nth_list_update)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   147
  done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   148
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   149
lemma merges_pres_le_ub:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   150
  "[| semilat(A,r,f); t:A; set ts <= A; set ss <= A; 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   151
     !p. p:set ps --> t <=_r ts!p; 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   152
     !p. p:set ps --> p<size ts; 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   153
     ss <=[r] ts |] 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   154
  ==> merges f t ps ss <=[r] ts"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   155
proof -
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   156
  { fix A r f t ts ps
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   157
    have
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   158
    "!!qs. [| semilat(A,r,f); set ts <= A; t:A; 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   159
       !p. p:set ps --> t <=_r ts!p; 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   160
       !p. p:set ps --> p<size ts |] ==> 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   161
    set qs <= set ps  --> 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   162
    (!ss. set ss <= A --> ss <=[r] ts --> merges f t qs ss <=[r] ts)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   163
    apply (induct_tac qs)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   164
     apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   165
    apply (simp (no_asm_simp))
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   166
    apply clarify
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   167
    apply (rotate_tac -2)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   168
    apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   169
    apply (erule allE, erule impE, erule_tac [2] mp)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   170
     apply (simp (no_asm_simp) add: closedD)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   171
    apply (simp add: list_update_le_listI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   172
    done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   173
  } note this [dest]
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   174
  
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   175
  case antecedent
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   176
  thus ?thesis by blast
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   177
qed
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   178
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   179
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   180
lemma nth_merges [rule_format]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   181
  "[| semilat (A, r, f); t:A; p < n |] ==> !ss. ss : list n A --> 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   182
     (!p:set ps. p<n) --> 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   183
     (merges f t ps ss)!p = (if p : set ps then t +_f ss!p else ss!p)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   184
  apply (induct_tac "ps")
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   185
   apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   186
  apply (simp add: nth_list_update closedD listE_nth_in) 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   187
  done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   188
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   189
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   190
(** propa **)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   191
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   192
lemma decomp_propa [rule_format]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   193
  "!ss w. (!q:set qs. q < size ss) --> 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   194
   propa f qs t ss w = 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   195
   (merges f t qs ss, {q. q:set qs & t +_f ss!q ~= ss!q} Un w)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   196
  apply (induct_tac qs)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   197
   apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   198
  apply (simp (no_asm))
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   199
  apply clarify
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   200
  apply (rule conjI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   201
   apply (simp add: nth_list_update)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   202
   apply blast
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   203
  apply (simp add: nth_list_update)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   204
  apply blast
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   205
  done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   206
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   207
(** iter **)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   208
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   209
ML_setup {*
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   210
let
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   211
val thy = the_context ()
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   212
val [iter_wf,iter_tc] = RecdefPackage.tcs_of thy "iter";
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   213
in
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   214
goalw_cterm [] (cterm_of (sign_of thy) (HOLogic.mk_Trueprop iter_wf));
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   215
by (REPEAT(rtac wf_same_fstI 1));
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   216
by (split_all_tac 1);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   217
by (asm_full_simp_tac (simpset() addsimps [thm "lesssub_def"]) 1);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   218
by (REPEAT(rtac wf_same_fstI 1));
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   219
by (rtac wf_lex_prod 1);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   220
 by (rtac wf_finite_psubset 2);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   221
by (Clarify_tac 1);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   222
by (dtac (thm "semilatDorderI" RS (thm "acc_le_listI")) 1);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   223
 by (assume_tac 1);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   224
by (rewrite_goals_tac [thm "acc_def",thm "lesssub_def"]);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   225
by (assume_tac 1);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   226
qed "iter_wf"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   227
end
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   228
*}
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   229
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   230
lemma inter_tc_lemma:  
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   231
  "[| semilat(A,r,f); ss : list n A; t:A; ! q:set qs. q < n; p:w |] ==> 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   232
      ss <[r] merges f t qs ss | 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   233
  merges f t qs ss = ss & {q. q:set qs & t +_f ss!q ~= ss!q} Un (w-{p}) < w"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   234
  apply (unfold lesssub_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   235
  apply (simp (no_asm_simp) add: merges_incr)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   236
  apply (rule impI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   237
  apply (rule merges_same_conv [THEN iffD1, elim_format]) 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   238
  apply assumption+
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   239
    defer
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   240
    apply (rule sym, assumption)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   241
   apply (simp cong add: conj_cong add: le_iff_plus_unchanged [symmetric])
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   242
   apply (blast intro!: psubsetI elim: equalityE)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   243
  apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   244
  done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   245
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   246
ML_setup {*
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   247
let
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   248
val thy = the_context ()
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   249
val [iter_wf,iter_tc] = RecdefPackage.tcs_of thy "iter";
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   250
in
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   251
goalw_cterm [] (cterm_of (sign_of thy) (HOLogic.mk_Trueprop iter_tc)); 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   252
by (asm_simp_tac (simpset() addsimps [same_fst_def,thm "pres_type_def"]) 1);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   253
by (clarify_tac (claset() delrules [disjCI]) 1);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   254
by (subgoal_tac "(@p. p:w) : w" 1);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   255
 by (fast_tac (claset() addIs [someI]) 2);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   256
by (subgoal_tac "ss : list (size ss) A" 1);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   257
 by (blast_tac (claset() addSIs [thm "listI"]) 2);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   258
by (subgoal_tac "!q:set(succs (@ p. p : w)). q < size ss" 1);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   259
 by (blast_tac (claset() addDs [thm "boundedD"]) 2);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   260
by (rotate_tac 1 1);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   261
by (asm_full_simp_tac (simpset() delsimps [thm "listE_length"]
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   262
      addsimps [thm "decomp_propa",finite_psubset_def,thm "inter_tc_lemma",
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   263
                bounded_nat_set_is_finite]) 1);
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   264
qed "iter_tc"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   265
end
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   266
*}
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   267
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   268
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   269
lemma iter_induct:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   270
  "(!!A r f step succs ss w. 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   271
    (!p. p = (@ p. p : w) & w ~= {} & semilat(A,r,f) & acc r & 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   272
      (!p:w. p < length ss) & bounded succs (length ss) & 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   273
         set ss <= A & pres_type step (length ss) A 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   274
         --> P A r f step succs (merges f (step p (ss!p)) (succs p) ss) 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   275
             ({q. q:set(succs p) & step p (ss!p) +_f ss!q ~= ss!q} Un (w-{p})))
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   276
  ==> P A r f step succs ss w) ==> P A r f step succs ss w"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   277
proof -
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   278
  case antecedent
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   279
  show ?thesis
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   280
    apply (rule_tac P = P in iter_tc [THEN iter_wf [THEN iter.induct]])
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   281
    apply (rule antecedent)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   282
    apply clarify
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   283
    apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   284
    apply (subgoal_tac "!q:set(succs (@ p. p : w)). q < size ss")
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   285
     apply (rotate_tac -1)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   286
     apply (simp add: decomp_propa)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   287
    apply (subgoal_tac "(@p. p:w) : w")
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   288
     apply (blast dest: boundedD)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   289
    apply (fast intro: someI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   290
    done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   291
qed
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   292
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   293
lemma iter_unfold:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   294
 "[| semilat(A,r,f); acc r; set ss <= A; pres_type step (length ss) A; 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   295
    bounded succs (size ss); !p:w. p<size ss |] ==> 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   296
 iter(((A,r,f),step,succs),ss,w) = 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   297
 (if w={} then ss 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   298
  else let p = SOME p. p : w 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   299
       in iter(((A,r,f),step,succs),merges f (step p (ss!p)) (succs p) ss, 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   300
        {q. q:set(succs p) & step p (ss!p) +_f ss!q ~= ss!q} Un (w-{p})))"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   301
  apply (rule iter_tc [THEN iter_wf [THEN iter.simps [THEN trans]]])
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   302
  apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   303
  apply (rule impI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   304
  apply (subst decomp_propa)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   305
   apply (subgoal_tac "(@p. p:w) : w")
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   306
    apply (blast dest: boundedD)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   307
   apply (fast intro: someI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   308
  apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   309
  done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   310
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   311
lemma stable_pres_lemma:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   312
  "[| semilat (A,r,f); pres_type step n A; bounded succs n; 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   313
     ss : list n A; p : w; ! q:w. q < n; 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   314
     ! q. q < n --> q ~: w --> stable r step succs ss q; q < n; 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   315
     q : set (succs p) --> step p (ss ! p) +_f ss ! q = ss ! q; 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   316
     q ~: w | q = p  |] 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   317
  ==> stable r step succs (merges f (step p (ss!p)) (succs p) ss) q"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   318
  apply (unfold stable_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   319
  apply (subgoal_tac "step p (ss!p) : A")
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   320
   defer
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   321
   apply (blast intro: listE_nth_in pres_typeD)     
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   322
  apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   323
  apply clarify
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   324
  apply (subst nth_merges)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   325
       apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   326
      apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   327
     prefer 2
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   328
     apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   329
    apply (blast dest: boundedD)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   330
   apply (blast dest: boundedD)  
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   331
  apply (subst nth_merges)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   332
       apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   333
      apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   334
     prefer 2
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   335
     apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   336
    apply (blast dest: boundedD)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   337
   apply (blast dest: boundedD)  
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   338
  apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   339
  apply (rule conjI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   340
   apply clarify
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   341
   apply (blast intro!: semilat_ub1 semilat_ub2 listE_nth_in
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   342
                intro: order_trans dest: boundedD)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   343
  apply (blast intro!: semilat_ub1 semilat_ub2 listE_nth_in
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   344
               intro: order_trans dest: boundedD)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   345
  done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   346
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   347
lemma merges_bounded_lemma [rule_format]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   348
  "[| semilat (A,r,f); mono r step n A; bounded succs n; 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   349
     step p (ss!p) : A; ss : list n A; ts : list n A; p < n; 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   350
     ss <=[r] ts; ! p. p < n --> stable r step succs ts p |] 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   351
  ==> merges f (step p (ss!p)) (succs p) ss <=[r] ts"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   352
  apply (unfold stable_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   353
  apply (blast intro!: listE_set monoD listE_nth_in le_listD less_lengthI
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   354
               intro: merges_pres_le_ub order_trans
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   355
               dest: pres_typeD boundedD)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   356
  done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   357
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   358
ML_setup {*
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   359
Unify.trace_bound := 80;
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   360
Unify.search_bound := 90;
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   361
*}
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   362
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   363
lemma iter_properties [rule_format]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   364
  "semilat(A,r,f) & acc r & pres_type step n A & mono r step n A & 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   365
  bounded succs n & (! p:w. p < n) & ss:list n A & 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   366
  (!p<n. p~:w --> stable r step succs ss p) 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   367
  --> iter(((A,r,f),step,succs),ss,w) : list n A & 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   368
      stables r step succs (iter(((A,r,f),step,succs),ss,w)) & 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   369
      ss <=[r] iter(((A,r,f),step,succs),ss,w) & 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   370
      (! ts:list n A. 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   371
           ss <=[r] ts & stables r step succs ts --> 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   372
           iter(((A,r,f),step,succs),ss,w) <=[r] ts)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   373
  apply (unfold stables_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   374
  apply (rule_tac A = A and r = r and f = f and step = step and ss = ss and w = w in iter_induct)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   375
  apply clarify
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   376
  apply (frule listE_length)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   377
  apply hypsubst
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   378
  apply (subst iter_unfold)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   379
        apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   380
       apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   381
      apply (simp (no_asm_simp))
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   382
     apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   383
    apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   384
   apply assumption
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   385
  apply (simp (no_asm_simp) del: listE_length)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   386
  apply (rule impI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   387
  apply (erule allE)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   388
  apply (erule impE)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   389
   apply (simp (no_asm_simp) del: listE_length)   
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   390
  apply (subgoal_tac "(@p. p:w) : w")
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   391
   prefer 2
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   392
   apply (fast intro: someI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   393
  apply (subgoal_tac "step (@ p. p : w) (ss ! (@ p. p : w)) : A")
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   394
   prefer 2
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   395
   apply (blast intro: pres_typeD listE_nth_in)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   396
  apply (erule impE)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   397
   apply (simp (no_asm_simp) del: listE_length le_iff_plus_unchanged [symmetric])
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   398
   apply (rule conjI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   399
    apply (blast dest: boundedD)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   400
   apply (rule conjI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   401
    apply (blast intro: merges_preserves_type dest: boundedD)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   402
   apply clarify
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   403
   apply (blast intro!: stable_pres_lemma)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   404
  apply (simp (no_asm_simp) del: listE_length)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   405
  apply (subgoal_tac "!q:set(succs (@ p. p : w)). q < size ss")
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   406
   prefer 2
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   407
   apply (blast dest: boundedD)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   408
  apply (subgoal_tac "step (@ p. p : w) (ss ! (@ p. p : w)) : A")
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   409
   prefer 2
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   410
   apply (blast intro: pres_typeD)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   411
  apply (rule conjI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   412
   apply (blast intro!: merges_incr intro: le_list_trans)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   413
  apply clarify
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   414
  apply (drule bspec, assumption, erule mp)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   415
  apply (simp del: listE_length)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   416
  apply (blast intro: merges_bounded_lemma)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   417
  done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   418
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   419
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   420
lemma is_dfa_kildall:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   421
  "[| semilat(A,r,f); acc r; pres_type step n A; 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   422
     mono r step n A; bounded succs n|] 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   423
  ==> is_dfa r (kildall (A,r,f) step succs) step succs n A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   424
  apply (unfold is_dfa_def kildall_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   425
  apply clarify
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   426
  apply simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   427
  apply (rule iter_properties)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   428
  apply (simp add: unstables_def stable_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   429
  apply (blast intro!: le_iff_plus_unchanged [THEN iffD2] listE_nth_in
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   430
               dest: boundedD pres_typeD)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   431
  done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   432
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   433
lemma is_bcv_kildall:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   434
  "[| semilat(A,r,f); acc r; top r T; 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   435
      pres_type step n A; bounded succs n; 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   436
      mono r step n A |]
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   437
  ==> is_bcv r T step succs n A (kildall (A,r,f) step succs)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   438
  apply (intro is_bcv_dfa is_dfa_kildall semilatDorderI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   439
  apply assumption+
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   440
  done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   441
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   442
end