src/HOL/MicroJava/DFA/Kildall.thy
author wenzelm
Fri, 18 Aug 2017 20:47:47 +0200
changeset 66453 cc19f7ca2ed6
parent 62042 6c6ccf573479
child 67443 3abf6a722518
permissions -rw-r--r--
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42150
b0c0638c4aad tuned headers;
wenzelm
parents: 41541
diff changeset
     1
(*  Title:      HOL/MicroJava/DFA/Kildall.thy
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     2
    Author:     Tobias Nipkow, Gerwin Klein
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     3
    Copyright   2000 TUM
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     4
*)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     5
61361
8b5f00202e1a isabelle update_cartouches;
wenzelm
parents: 58886
diff changeset
     6
section \<open>Kildall's Algorithm \label{sec:Kildall}\<close>
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     7
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     8
theory Kildall
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 62042
diff changeset
     9
imports SemilatAlg "HOL-Library.While_Combinator"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    10
begin
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    11
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    12
primrec propa :: "'s binop \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's list \<Rightarrow> nat set \<Rightarrow> 's list * nat set" where
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    13
  "propa f []      ss w = (ss,w)"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    14
| "propa f (q'#qs) ss w = (let (q,t) = q';
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    15
                               u = t +_f ss!q;
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    16
                               w' = (if u = ss!q then w else insert q w)
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    17
                           in propa f qs (ss[q := u]) w')"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    18
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    19
definition iter :: "'s binop \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> nat set \<Rightarrow> 's list \<times> nat set" where
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    20
  "iter f step ss w == while (%(ss,w). w \<noteq> {})
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    21
       (%(ss,w). let p = SOME p. p \<in> w
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    22
                 in propa f (step p (ss!p)) ss (w-{p}))
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    23
       (ss,w)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    24
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    25
definition unstables :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> nat set" where
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    26
"unstables r step ss == {p. p < size ss \<and> \<not>stable r step ss p}"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    27
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    28
definition kildall :: "'s ord \<Rightarrow> 's binop \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> 's list" where
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    29
"kildall r f step ss == fst(iter f step ss (unstables r step ss))"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    30
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    31
primrec merges :: "'s binop \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's list \<Rightarrow> 's list" where
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    32
  "merges f []      ss = ss"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    33
| "merges f (p'#ps) ss = (let (p,s) = p' in merges f ps (ss[p := s +_f ss!p]))"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    34
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    35
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    36
lemmas [simp] = Let_def Semilat.le_iff_plus_unchanged [OF Semilat.intro, symmetric]
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    37
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    38
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    39
lemma (in Semilat) nth_merges:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    40
 "\<And>ss. \<lbrakk>p < length ss; ss \<in> list n A; \<forall>(p,t)\<in>set ps. p<n \<and> t\<in>A \<rbrakk> \<Longrightarrow>
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    41
  (merges f ps ss)!p = map snd [(p',t') \<leftarrow> ps. p'=p] ++_f ss!p"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    42
  (is "\<And>ss. \<lbrakk>_; _; ?steptype ps\<rbrakk> \<Longrightarrow> ?P ss ps")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    43
proof (induct ps)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    44
  show "\<And>ss. ?P ss []" by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    45
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    46
  fix ss p' ps'
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    47
  assume ss: "ss \<in> list n A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    48
  assume l:  "p < length ss"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    49
  assume "?steptype (p'#ps')"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    50
  then obtain a b where
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    51
    p': "p'=(a,b)" and ab: "a<n" "b\<in>A" and ps': "?steptype ps'"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    52
    by (cases p') auto
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    53
  assume "\<And>ss. p< length ss \<Longrightarrow> ss \<in> list n A \<Longrightarrow> ?steptype ps' \<Longrightarrow> ?P ss ps'"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    54
  from this [OF _ _ ps'] have IH: "\<And>ss. ss \<in> list n A \<Longrightarrow> p < length ss \<Longrightarrow> ?P ss ps'" .
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    55
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    56
  from ss ab
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    57
  have "ss[a := b +_f ss!a] \<in> list n A" by (simp add: closedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    58
  moreover
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    59
  from calculation l
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    60
  have "p < length (ss[a := b +_f ss!a])" by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    61
  ultimately
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    62
  have "?P (ss[a := b +_f ss!a]) ps'" by (rule IH)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    63
  with p' l
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    64
  show "?P ss (p'#ps')" by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    65
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    66
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    67
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    68
(** merges **)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    69
52847
820339715ffe tuned proofs;
wenzelm
parents: 46226
diff changeset
    70
lemma length_merges [simp]: "size(merges f ps ss) = size ss"
820339715ffe tuned proofs;
wenzelm
parents: 46226
diff changeset
    71
  by (induct ps arbitrary: ss) auto
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    72
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    73
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    74
lemma (in Semilat) merges_preserves_type_lemma:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    75
shows "\<forall>xs. xs \<in> list n A \<longrightarrow> (\<forall>(p,x) \<in> set ps. p<n \<and> x\<in>A)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    76
          \<longrightarrow> merges f ps xs \<in> list n A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    77
apply (insert closedI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    78
apply (unfold closed_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    79
apply (induct_tac ps)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    80
 apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    81
apply clarsimp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    82
done
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    83
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    84
lemma (in Semilat) merges_preserves_type [simp]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    85
 "\<lbrakk> xs \<in> list n A; \<forall>(p,x) \<in> set ps. p<n \<and> x\<in>A \<rbrakk>
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    86
  \<Longrightarrow> merges f ps xs \<in> list n A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    87
by (simp add: merges_preserves_type_lemma)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    88
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    89
lemma (in Semilat) merges_incr_lemma:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    90
 "\<forall>xs. xs \<in> list n A \<longrightarrow> (\<forall>(p,x)\<in>set ps. p<size xs \<and> x \<in> A) \<longrightarrow> xs <=[r] merges f ps xs"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    91
apply (induct_tac ps)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    92
 apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    93
apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    94
apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    95
apply (rule order_trans)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    96
  apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    97
 apply (erule list_update_incr)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    98
  apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    99
 apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   100
apply (blast intro!: listE_set intro: closedD listE_length [THEN nth_in])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   101
done
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   102
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   103
lemma (in Semilat) merges_incr:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   104
 "\<lbrakk> xs \<in> list n A; \<forall>(p,x)\<in>set ps. p<size xs \<and> x \<in> A \<rbrakk> 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   105
  \<Longrightarrow> xs <=[r] merges f ps xs"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   106
  by (simp add: merges_incr_lemma)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   107
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   108
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   109
lemma (in Semilat) merges_same_conv [rule_format]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   110
 "(\<forall>xs. xs \<in> list n A \<longrightarrow> (\<forall>(p,x)\<in>set ps. p<size xs \<and> x\<in>A) \<longrightarrow> 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   111
     (merges f ps xs = xs) = (\<forall>(p,x)\<in>set ps. x <=_r xs!p))"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   112
  apply (induct_tac ps)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   113
   apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   114
  apply clarsimp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   115
  apply (rename_tac p x ps xs)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   116
  apply (rule iffI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   117
   apply (rule context_conjI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   118
    apply (subgoal_tac "xs[p := x +_f xs!p] <=[r] xs")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   119
     apply (drule_tac p = p in le_listD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   120
      apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   121
     apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   122
    apply (erule subst, rule merges_incr)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   123
       apply (blast intro!: listE_set intro: closedD listE_length [THEN nth_in])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   124
      apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   125
      apply (rule conjI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   126
       apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   127
       apply (blast dest: boundedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   128
      apply blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   129
   apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   130
   apply (erule allE)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   131
   apply (erule impE)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   132
    apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   133
   apply (drule bspec)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   134
    apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   135
   apply (simp add: le_iff_plus_unchanged [THEN iffD1] list_update_same_conv [THEN iffD2])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   136
   apply blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   137
  apply clarify 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   138
  apply (simp add: le_iff_plus_unchanged [THEN iffD1] list_update_same_conv [THEN iffD2])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   139
  done
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   140
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   141
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   142
lemma (in Semilat) list_update_le_listI [rule_format]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   143
  "set xs <= A \<longrightarrow> set ys <= A \<longrightarrow> xs <=[r] ys \<longrightarrow> p < size xs \<longrightarrow>  
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   144
   x <=_r ys!p \<longrightarrow> x\<in>A \<longrightarrow> xs[p := x +_f xs!p] <=[r] ys"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   145
  apply(insert semilat)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   146
  apply (unfold Listn.le_def lesub_def semilat_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   147
  apply (simp add: list_all2_conv_all_nth nth_list_update)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   148
  done
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   149
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   150
lemma (in Semilat) merges_pres_le_ub:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   151
  assumes "set ts <= A" and "set ss <= A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   152
    and "\<forall>(p,t)\<in>set ps. t <=_r ts!p \<and> t \<in> A \<and> p < size ts" and "ss <=[r] ts"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   153
  shows "merges f ps ss <=[r] ts"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   154
proof -
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   155
  { fix t ts ps
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   156
    have
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   157
    "\<And>qs. \<lbrakk>set ts <= A; \<forall>(p,t)\<in>set ps. t <=_r ts!p \<and> t \<in> A \<and> p< size ts \<rbrakk> \<Longrightarrow>
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   158
    set qs <= set ps  \<longrightarrow> 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   159
    (\<forall>ss. set ss <= A \<longrightarrow> ss <=[r] ts \<longrightarrow> merges f qs ss <=[r] ts)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   160
    apply (induct_tac qs)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   161
     apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   162
    apply (simp (no_asm_simp))
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   163
    apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   164
    apply (rotate_tac -2)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   165
    apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   166
    apply (erule allE, erule impE, erule_tac [2] mp)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   167
     apply (drule bspec, assumption)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   168
     apply (simp add: closedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   169
    apply (drule bspec, assumption)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   170
    apply (simp add: list_update_le_listI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   171
    done 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   172
  } note this [dest]
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   173
  
41541
1fa4725c4656 eliminated global prems;
wenzelm
parents: 41413
diff changeset
   174
  from assms show ?thesis by blast
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   175
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   176
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   177
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   178
(** propa **)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   179
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   180
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   181
lemma decomp_propa:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   182
  "\<And>ss w. (\<forall>(q,t)\<in>set qs. q < size ss) \<Longrightarrow> 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   183
   propa f qs ss w = 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   184
   (merges f qs ss, {q. \<exists>t. (q,t)\<in>set qs \<and> t +_f ss!q \<noteq> ss!q} Un w)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   185
  apply (induct qs)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   186
   apply simp   
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   187
  apply (simp (no_asm))
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   188
  apply clarify  
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   189
  apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   190
  apply (rule conjI) 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   191
   apply blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   192
  apply (simp add: nth_list_update)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   193
  apply blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   194
  done 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   195
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   196
(** iter **)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   197
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   198
lemma (in Semilat) stable_pres_lemma:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   199
shows "\<lbrakk>pres_type step n A; bounded step n; 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   200
     ss \<in> list n A; p \<in> w; \<forall>q\<in>w. q < n; 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   201
     \<forall>q. q < n \<longrightarrow> q \<notin> w \<longrightarrow> stable r step ss q; q < n; 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   202
     \<forall>s'. (q,s') \<in> set (step p (ss ! p)) \<longrightarrow> s' +_f ss ! q = ss ! q; 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   203
     q \<notin> w \<or> q = p \<rbrakk> 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   204
  \<Longrightarrow> stable r step (merges f (step p (ss!p)) ss) q"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   205
  apply (unfold stable_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   206
  apply (subgoal_tac "\<forall>s'. (q,s') \<in> set (step p (ss!p)) \<longrightarrow> s' : A")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   207
   prefer 2
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   208
   apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   209
   apply (erule pres_typeD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   210
    prefer 3 apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   211
    apply (rule listE_nth_in)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   212
     apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   213
    apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   214
   apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   215
  apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   216
  apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   217
  apply (subst nth_merges)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   218
       apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   219
       apply (blast dest: boundedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   220
      apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   221
     apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   222
     apply (rule conjI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   223
      apply (blast dest: boundedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   224
     apply (erule pres_typeD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   225
       prefer 3 apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   226
      apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   227
     apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   228
apply(subgoal_tac "q < length ss")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   229
prefer 2 apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   230
  apply (frule nth_merges [of q _ _ "step p (ss!p)"]) (* fixme: why does method subst not work?? *)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   231
apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   232
  apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   233
  apply (rule conjI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   234
   apply (blast dest: boundedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   235
  apply (erule pres_typeD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   236
     prefer 3 apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   237
    apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   238
   apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   239
  apply (drule_tac P = "\<lambda>x. (a, b) \<in> set (step q x)" in subst)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   240
   apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   241
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   242
 apply (simp add: plusplus_empty)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   243
 apply (cases "q \<in> w")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   244
  apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   245
  apply (rule ub1')
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   246
     apply (rule semilat)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   247
    apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   248
    apply (rule pres_typeD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   249
       apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   250
      prefer 3 apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   251
     apply (blast intro: listE_nth_in dest: boundedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   252
    apply (blast intro: pres_typeD dest: boundedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   253
   apply (blast intro: listE_nth_in dest: boundedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   254
  apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   255
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   256
 apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   257
 apply (erule allE, erule impE, assumption, erule impE, assumption)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   258
 apply (rule order_trans)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   259
   apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   260
  defer
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   261
 apply (rule pp_ub2)(*
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   262
    apply assumption*)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   263
   apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   264
   apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   265
   apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   266
   apply (rule pres_typeD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   267
      apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   268
     prefer 3 apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   269
    apply (blast intro: listE_nth_in dest: boundedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   270
   apply (blast intro: pres_typeD dest: boundedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   271
  apply (blast intro: listE_nth_in dest: boundedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   272
 apply blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   273
 done
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   274
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   275
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   276
lemma (in Semilat) merges_bounded_lemma:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   277
 "\<lbrakk> mono r step n A; bounded step n; 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   278
    \<forall>(p',s') \<in> set (step p (ss!p)). s' \<in> A; ss \<in> list n A; ts \<in> list n A; p < n; 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   279
    ss <=[r] ts; \<forall>p. p < n \<longrightarrow> stable r step ts p \<rbrakk> 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   280
  \<Longrightarrow> merges f (step p (ss!p)) ss <=[r] ts" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   281
  apply (unfold stable_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   282
  apply (rule merges_pres_le_ub)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   283
     apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   284
    apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   285
   prefer 2 apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   286
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   287
  apply clarsimp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   288
  apply (drule boundedD, assumption+)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   289
  apply (erule allE, erule impE, assumption)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   290
  apply (drule bspec, assumption)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   291
  apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   292
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   293
  apply (drule monoD [of _ _ _ _ p "ss!p"  "ts!p"])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   294
     apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   295
    apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   296
   apply (simp add: le_listD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   297
  
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   298
  apply (drule lesub_step_typeD, assumption) 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   299
  apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   300
  apply (drule bspec, assumption)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   301
  apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   302
  apply (blast intro: order_trans)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   303
  done
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   304
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   305
lemma termination_lemma:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   306
  assumes semilat: "semilat (A, r, f)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   307
  shows "\<lbrakk> ss \<in> list n A; \<forall>(q,t)\<in>set qs. q<n \<and> t\<in>A; p\<in>w \<rbrakk> \<Longrightarrow> 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   308
  ss <[r] merges f qs ss \<or> 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   309
  merges f qs ss = ss \<and> {q. \<exists>t. (q,t)\<in>set qs \<and> t +_f ss!q \<noteq> ss!q} Un (w-{p}) < w" (is "PROP ?P")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   310
proof -
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   311
  interpret Semilat A r f using assms by (rule Semilat.intro)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   312
  show "PROP ?P" apply(insert semilat)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   313
    apply (unfold lesssub_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   314
    apply (simp (no_asm_simp) add: merges_incr)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   315
    apply (rule impI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   316
    apply (rule merges_same_conv [THEN iffD1, elim_format]) 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   317
    apply assumption+
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   318
    defer
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   319
    apply (rule sym, assumption)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   320
    defer apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   321
    apply (subgoal_tac "\<forall>q t. \<not>((q, t) \<in> set qs \<and> t +_f ss ! q \<noteq> ss ! q)")
46226
e88e980ed735 tuned proofs;
wenzelm
parents: 42150
diff changeset
   322
    apply (blast elim: equalityE)
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   323
    apply clarsimp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   324
    apply (drule bspec, assumption) 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   325
    apply (drule bspec, assumption)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   326
    apply clarsimp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   327
    done
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   328
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   329
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   330
lemma iter_properties[rule_format]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   331
  assumes semilat: "semilat (A, r, f)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   332
  shows "\<lbrakk> acc r ; pres_type step n A; mono r step n A;
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   333
     bounded step n; \<forall>p\<in>w0. p < n; ss0 \<in> list n A;
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   334
     \<forall>p<n. p \<notin> w0 \<longrightarrow> stable r step ss0 p \<rbrakk> \<Longrightarrow>
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   335
   iter f step ss0 w0 = (ss',w')
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   336
   \<longrightarrow>
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   337
   ss' \<in> list n A \<and> stables r step ss' \<and> ss0 <=[r] ss' \<and>
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   338
   (\<forall>ts\<in>list n A. ss0 <=[r] ts \<and> stables r step ts \<longrightarrow> ss' <=[r] ts)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   339
  (is "PROP ?P")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   340
proof -
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   341
  interpret Semilat A r f using assms by (rule Semilat.intro)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   342
  show "PROP ?P" apply(insert semilat)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   343
apply (unfold iter_def stables_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   344
apply (rule_tac P = "%(ss,w).
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   345
 ss \<in> list n A \<and> (\<forall>p<n. p \<notin> w \<longrightarrow> stable r step ss p) \<and> ss0 <=[r] ss \<and>
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   346
 (\<forall>ts\<in>list n A. ss0 <=[r] ts \<and> stables r step ts \<longrightarrow> ss <=[r] ts) \<and>
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   347
 (\<forall>p\<in>w. p < n)" and
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   348
 r = "{(ss',ss) . ss <[r] ss'} <*lex*> finite_psubset"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   349
       in while_rule)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   350
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 61361
diff changeset
   351
\<comment> "Invariant holds initially:"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   352
apply (simp add:stables_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   353
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 61361
diff changeset
   354
\<comment> "Invariant is preserved:"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   355
apply(simp add: stables_def split_paired_all)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   356
apply(rename_tac ss w)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   357
apply(subgoal_tac "(SOME p. p \<in> w) \<in> w")
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 52847
diff changeset
   358
 prefer 2 apply (fast intro: someI)
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   359
apply(subgoal_tac "\<forall>(q,t) \<in> set (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))). q < length ss \<and> t \<in> A")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   360
 prefer 2
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   361
 apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   362
 apply (rule conjI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   363
  apply(clarsimp, blast dest!: boundedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   364
 apply (erule pres_typeD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   365
  prefer 3
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   366
  apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   367
  apply (erule listE_nth_in)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   368
  apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   369
 apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   370
apply (subst decomp_propa)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   371
 apply fast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   372
apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   373
apply (rule conjI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   374
 apply (rule merges_preserves_type)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   375
 apply blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   376
 apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   377
 apply (rule conjI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   378
  apply(clarsimp, fast dest!: boundedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   379
 apply (erule pres_typeD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   380
  prefer 3
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   381
  apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   382
  apply (erule listE_nth_in)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   383
  apply blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   384
 apply blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   385
apply (rule conjI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   386
 apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   387
 apply (blast intro!: stable_pres_lemma)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   388
apply (rule conjI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   389
 apply (blast intro!: merges_incr intro: le_list_trans)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   390
apply (rule conjI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   391
 apply clarsimp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   392
 apply (blast intro!: merges_bounded_lemma)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   393
apply (blast dest!: boundedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   394
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   395
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 61361
diff changeset
   396
\<comment> "Postcondition holds upon termination:"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   397
apply(clarsimp simp add: stables_def split_paired_all)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   398
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 61361
diff changeset
   399
\<comment> "Well-foundedness of the termination relation:"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   400
apply (rule wf_lex_prod)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   401
 apply (insert orderI [THEN acc_le_listI])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   402
 apply (simp add: acc_def lesssub_def wfP_wf_eq [symmetric])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   403
apply (rule wf_finite_psubset) 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   404
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 61361
diff changeset
   405
\<comment> "Loop decreases along termination relation:"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   406
apply(simp add: stables_def split_paired_all)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   407
apply(rename_tac ss w)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   408
apply(subgoal_tac "(SOME p. p \<in> w) \<in> w")
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 52847
diff changeset
   409
 prefer 2 apply (fast intro: someI)
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   410
apply(subgoal_tac "\<forall>(q,t) \<in> set (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))). q < length ss \<and> t \<in> A")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   411
 prefer 2
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   412
 apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   413
 apply (rule conjI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   414
  apply(clarsimp, blast dest!: boundedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   415
 apply (erule pres_typeD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   416
  prefer 3
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   417
  apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   418
  apply (erule listE_nth_in)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   419
  apply blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   420
 apply blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   421
apply (subst decomp_propa)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   422
 apply blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   423
apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   424
apply (simp del: listE_length
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   425
    add: lex_prod_def finite_psubset_def 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   426
         bounded_nat_set_is_finite)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   427
apply (rule termination_lemma)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   428
apply assumption+
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   429
defer
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   430
apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   431
apply clarsimp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   432
done
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   433
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   434
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   435
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   436
lemma kildall_properties:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   437
assumes semilat: "semilat (A, r, f)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   438
shows "\<lbrakk> acc r; pres_type step n A; mono r step n A;
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   439
     bounded step n; ss0 \<in> list n A \<rbrakk> \<Longrightarrow>
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   440
  kildall r f step ss0 \<in> list n A \<and>
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   441
  stables r step (kildall r f step ss0) \<and>
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   442
  ss0 <=[r] kildall r f step ss0 \<and>
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   443
  (\<forall>ts\<in>list n A. ss0 <=[r] ts \<and> stables r step ts \<longrightarrow>
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   444
                 kildall r f step ss0 <=[r] ts)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   445
  (is "PROP ?P")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   446
proof -
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   447
  interpret Semilat A r f using assms by (rule Semilat.intro)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   448
  show "PROP ?P"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   449
apply (unfold kildall_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   450
apply(case_tac "iter f step ss0 (unstables r step ss0)")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   451
apply(simp)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   452
apply (rule iter_properties)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   453
apply (simp_all add: unstables_def stable_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   454
apply (rule semilat)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   455
done
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   456
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   457
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   458
lemma is_bcv_kildall:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   459
assumes semilat: "semilat (A, r, f)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   460
shows "\<lbrakk> acc r; top r T; pres_type step n A; bounded step n; mono r step n A \<rbrakk>
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   461
  \<Longrightarrow> is_bcv r T step n A (kildall r f step)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   462
  (is "PROP ?P")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   463
proof -
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   464
  interpret Semilat A r f using assms by (rule Semilat.intro)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   465
  show "PROP ?P"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   466
apply(unfold is_bcv_def wt_step_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   467
apply(insert semilat kildall_properties[of A])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   468
apply(simp add:stables_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   469
apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   470
apply(subgoal_tac "kildall r f step ss \<in> list n A")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   471
 prefer 2 apply (simp(no_asm_simp))
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   472
apply (rule iffI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   473
 apply (rule_tac x = "kildall r f step ss" in bexI) 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   474
  apply (rule conjI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   475
   apply (blast)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   476
  apply (simp  (no_asm_simp))
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   477
 apply(assumption)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   478
apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   479
apply(subgoal_tac "kildall r f step ss!p <=_r ts!p")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   480
 apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   481
apply (blast intro!: le_listD less_lengthI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   482
done
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   483
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   484
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   485
end