src/HOL/MicroJava/DFA/Err.thy
author wenzelm
Tue, 05 Jan 2010 13:45:17 +0100
changeset 34262 a6d2d9c07e46
parent 33954 1bc3b688548c
child 35102 cc7a0b9f938c
permissions -rwxr-xr-x
added Promise.fulfill_result;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/BV/Err.thy
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     2
    Author:     Tobias Nipkow
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
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     6
header {* \isaheader{The Error Type} *}
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 Err
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     9
imports Semilat
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
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    12
datatype 'a err = Err | OK 'a
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    13
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    14
types 'a ebinop = "'a \<Rightarrow> 'a \<Rightarrow> 'a err"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    15
      'a esl =    "'a set * 'a ord * 'a ebinop"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    16
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    17
consts
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    18
  ok_val :: "'a err \<Rightarrow> 'a"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    19
primrec
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    20
  "ok_val (OK x) = x"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    21
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    22
constdefs
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    23
 lift :: "('a \<Rightarrow> 'b err) \<Rightarrow> ('a err \<Rightarrow> 'b err)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    24
"lift f e == case e of Err \<Rightarrow> Err | OK x \<Rightarrow> f x"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    25
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    26
 lift2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c err) \<Rightarrow> 'a err \<Rightarrow> 'b err \<Rightarrow> 'c err"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    27
"lift2 f e1 e2 ==
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    28
 case e1 of Err  \<Rightarrow> Err
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    29
          | OK x \<Rightarrow> (case e2 of Err \<Rightarrow> Err | OK y \<Rightarrow> f x y)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    30
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    31
 le :: "'a ord \<Rightarrow> 'a err ord"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    32
"le r e1 e2 ==
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    33
        case e2 of Err \<Rightarrow> True |
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    34
                   OK y \<Rightarrow> (case e1 of Err \<Rightarrow> False | OK x \<Rightarrow> x <=_r y)"
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
 sup :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a err \<Rightarrow> 'b err \<Rightarrow> 'c err)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    37
"sup f == lift2(%x y. OK(x +_f y))"
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
 err :: "'a set \<Rightarrow> 'a err set"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    40
"err A == insert Err {x . ? y:A. x = OK y}"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    41
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    42
 esl :: "'a sl \<Rightarrow> 'a esl"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    43
"esl == %(A,r,f). (A,r, %x y. OK(f x y))"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    44
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    45
 sl :: "'a esl \<Rightarrow> 'a err sl"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    46
"sl == %(A,r,f). (err A, le r, lift2 f)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    47
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    48
syntax
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    49
 err_semilat :: "'a esl \<Rightarrow> bool"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    50
translations
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    51
"err_semilat L" == "semilat(Err.sl L)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    52
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    53
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    54
consts
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    55
  strict  :: "('a \<Rightarrow> 'b err) \<Rightarrow> ('a err \<Rightarrow> 'b err)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    56
primrec
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    57
  "strict f Err    = Err"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    58
  "strict f (OK x) = f x"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    59
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    60
lemma strict_Some [simp]: 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    61
  "(strict f x = OK y) = (\<exists> z. x = OK z \<and> f z = OK y)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    62
  by (cases x, auto)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    63
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    64
lemma not_Err_eq:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    65
  "(x \<noteq> Err) = (\<exists>a. x = OK a)" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    66
  by (cases x) auto
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
lemma not_OK_eq:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    69
  "(\<forall>y. x \<noteq> OK y) = (x = Err)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    70
  by (cases x) auto  
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    71
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    72
lemma unfold_lesub_err:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    73
  "e1 <=_(le r) e2 == le r e1 e2"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    74
  by (simp add: lesub_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    75
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    76
lemma le_err_refl:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    77
  "!x. x <=_r x \<Longrightarrow> e <=_(Err.le r) e"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    78
apply (unfold lesub_def Err.le_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    79
apply (simp split: err.split)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    80
done 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    81
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    82
lemma le_err_trans [rule_format]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    83
  "order r \<Longrightarrow> e1 <=_(le r) e2 \<longrightarrow> e2 <=_(le r) e3 \<longrightarrow> e1 <=_(le r) e3"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    84
apply (unfold unfold_lesub_err le_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    85
apply (simp split: err.split)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    86
apply (blast intro: order_trans)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    87
done
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 le_err_antisym [rule_format]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    90
  "order r \<Longrightarrow> e1 <=_(le r) e2 \<longrightarrow> e2 <=_(le r) e1 \<longrightarrow> e1=e2"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    91
apply (unfold unfold_lesub_err le_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    92
apply (simp split: err.split)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    93
apply (blast intro: order_antisym)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    94
done 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    95
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    96
lemma OK_le_err_OK:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    97
  "(OK x <=_(le r) OK y) = (x <=_r y)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    98
  by (simp add: unfold_lesub_err le_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    99
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   100
lemma order_le_err [iff]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   101
  "order(le r) = order r"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   102
apply (rule iffI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   103
 apply (subst Semilat.order_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   104
 apply (blast dest: order_antisym OK_le_err_OK [THEN iffD2]
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   105
              intro: order_trans OK_le_err_OK [THEN iffD1])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   106
apply (subst Semilat.order_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   107
apply (blast intro: le_err_refl le_err_trans le_err_antisym
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   108
             dest: order_refl)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   109
done 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   110
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   111
lemma le_Err [iff]:  "e <=_(le r) Err"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   112
  by (simp add: unfold_lesub_err le_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   113
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   114
lemma Err_le_conv [iff]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   115
 "Err <=_(le r) e  = (e = Err)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   116
  by (simp add: unfold_lesub_err le_def  split: err.split)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   117
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   118
lemma le_OK_conv [iff]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   119
  "e <=_(le r) OK x  =  (? y. e = OK y & y <=_r x)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   120
  by (simp add: unfold_lesub_err le_def split: err.split)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   121
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   122
lemma OK_le_conv:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   123
 "OK x <=_(le r) e  =  (e = Err | (? y. e = OK y & x <=_r y))"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   124
  by (simp add: unfold_lesub_err le_def split: err.split)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   125
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   126
lemma top_Err [iff]: "top (le r) Err";
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   127
  by (simp add: top_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   128
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   129
lemma OK_less_conv [rule_format, iff]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   130
  "OK x <_(le r) e = (e=Err | (? y. e = OK y & x <_r y))"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   131
  by (simp add: lesssub_def lesub_def le_def split: err.split)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   132
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   133
lemma not_Err_less [rule_format, iff]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   134
  "~(Err <_(le r) x)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   135
  by (simp add: lesssub_def lesub_def le_def split: err.split)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   136
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   137
lemma semilat_errI [intro]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   138
  assumes semilat: "semilat (A, r, f)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   139
  shows "semilat(err A, Err.le r, lift2(%x y. OK(f x y)))"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   140
  apply(insert semilat)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   141
  apply (unfold semilat_Def closed_def plussub_def lesub_def 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   142
    lift2_def Err.le_def err_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   143
  apply (simp split: err.split)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   144
  done
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   145
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   146
lemma err_semilat_eslI_aux:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   147
  assumes semilat: "semilat (A, r, f)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   148
  shows "err_semilat(esl(A,r,f))"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   149
  apply (unfold sl_def esl_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   150
  apply (simp add: semilat_errI[OF semilat])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   151
  done
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   152
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   153
lemma err_semilat_eslI [intro, simp]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   154
 "\<And>L. semilat L \<Longrightarrow> err_semilat(esl L)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   155
by(simp add: err_semilat_eslI_aux split_tupled_all)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   156
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   157
lemma acc_err [simp, intro!]:  "acc r \<Longrightarrow> acc(le r)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   158
apply (unfold acc_def lesub_def le_def lesssub_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   159
apply (simp add: wf_eq_minimal split: err.split)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   160
apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   161
apply (case_tac "Err : Q")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   162
 apply blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   163
apply (erule_tac x = "{a . OK a : Q}" in allE)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   164
apply (case_tac "x")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   165
 apply fast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   166
apply blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   167
done 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   168
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   169
lemma Err_in_err [iff]: "Err : err A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   170
  by (simp add: err_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   171
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   172
lemma Ok_in_err [iff]: "(OK x : err A) = (x:A)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   173
  by (auto simp add: err_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   174
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   175
section {* lift *}
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
lemma lift_in_errI:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   178
  "\<lbrakk> e : err S; !x:S. e = OK x \<longrightarrow> f x : err S \<rbrakk> \<Longrightarrow> lift f e : err S"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   179
apply (unfold lift_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   180
apply (simp split: err.split)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   181
apply blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   182
done 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   183
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   184
lemma Err_lift2 [simp]: 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   185
  "Err +_(lift2 f) x = Err"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   186
  by (simp add: lift2_def plussub_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   187
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   188
lemma lift2_Err [simp]: 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   189
  "x +_(lift2 f) Err = Err"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   190
  by (simp add: lift2_def plussub_def split: err.split)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   191
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   192
lemma OK_lift2_OK [simp]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   193
  "OK x +_(lift2 f) OK y = x +_f y"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   194
  by (simp add: lift2_def plussub_def split: err.split)
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
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   197
section {* sup *}
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   198
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   199
lemma Err_sup_Err [simp]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   200
  "Err +_(Err.sup f) x = Err"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   201
  by (simp add: plussub_def Err.sup_def Err.lift2_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   202
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   203
lemma Err_sup_Err2 [simp]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   204
  "x +_(Err.sup f) Err = Err"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   205
  by (simp add: plussub_def Err.sup_def Err.lift2_def split: err.split)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   206
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   207
lemma Err_sup_OK [simp]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   208
  "OK x +_(Err.sup f) OK y = OK(x +_f y)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   209
  by (simp add: plussub_def Err.sup_def Err.lift2_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   210
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   211
lemma Err_sup_eq_OK_conv [iff]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   212
  "(Err.sup f ex ey = OK z) = (? x y. ex = OK x & ey = OK y & f x y = z)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   213
apply (unfold Err.sup_def lift2_def plussub_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   214
apply (rule iffI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   215
 apply (simp split: err.split_asm)
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 simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   218
done
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   219
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   220
lemma Err_sup_eq_Err [iff]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   221
  "(Err.sup f ex ey = Err) = (ex=Err | ey=Err)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   222
apply (unfold Err.sup_def lift2_def plussub_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   223
apply (simp split: err.split)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   224
done 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   225
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   226
section {* semilat (err A) (le r) f *}
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   227
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   228
lemma semilat_le_err_Err_plus [simp]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   229
  "\<lbrakk> x: err A; semilat(err A, le r, f) \<rbrakk> \<Longrightarrow> Err +_f x = Err"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   230
  by (blast intro: Semilat.le_iff_plus_unchanged [OF Semilat.intro, THEN iffD1]
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   231
                   Semilat.le_iff_plus_unchanged2 [OF Semilat.intro, THEN iffD1])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   232
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   233
lemma semilat_le_err_plus_Err [simp]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   234
  "\<lbrakk> x: err A; semilat(err A, le r, f) \<rbrakk> \<Longrightarrow> x +_f Err = Err"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   235
  by (blast intro: Semilat.le_iff_plus_unchanged [OF Semilat.intro, THEN iffD1]
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   236
                   Semilat.le_iff_plus_unchanged2 [OF Semilat.intro, THEN iffD1])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   237
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   238
lemma semilat_le_err_OK1:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   239
  "\<lbrakk> x:A; y:A; semilat(err A, le r, f); OK x +_f OK y = OK z \<rbrakk> 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   240
  \<Longrightarrow> x <=_r z";
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   241
apply (rule OK_le_err_OK [THEN iffD1])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   242
apply (erule subst)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   243
apply (simp add: Semilat.ub1 [OF Semilat.intro])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   244
done
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   245
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   246
lemma semilat_le_err_OK2:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   247
  "\<lbrakk> x:A; y:A; semilat(err A, le r, f); OK x +_f OK y = OK z \<rbrakk> 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   248
  \<Longrightarrow> y <=_r z"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   249
apply (rule OK_le_err_OK [THEN iffD1])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   250
apply (erule subst)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   251
apply (simp add: Semilat.ub2 [OF Semilat.intro])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   252
done
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   253
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   254
lemma eq_order_le:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   255
  "\<lbrakk> x=y; order r \<rbrakk> \<Longrightarrow> x <=_r y"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   256
apply (unfold Semilat.order_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   257
apply blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   258
done
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   259
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   260
lemma OK_plus_OK_eq_Err_conv [simp]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   261
  assumes "x:A" and "y:A" and "semilat(err A, le r, fe)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   262
  shows "((OK x) +_fe (OK y) = Err) = (~(? z:A. x <=_r z & y <=_r z))"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   263
proof -
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   264
  have plus_le_conv3: "\<And>A x y z f r. 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   265
    \<lbrakk> semilat (A,r,f); x +_f y <=_r z; x:A; y:A; z:A \<rbrakk> 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   266
    \<Longrightarrow> x <=_r z \<and> y <=_r z"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   267
    by (rule Semilat.plus_le_conv [OF Semilat.intro, THEN iffD1])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   268
  from prems show ?thesis
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   269
  apply (rule_tac iffI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   270
   apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   271
   apply (drule OK_le_err_OK [THEN iffD2])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   272
   apply (drule OK_le_err_OK [THEN iffD2])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   273
   apply (drule Semilat.lub [OF Semilat.intro, of _ _ _ "OK x" _ "OK y"])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   274
        apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   275
       apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   276
      apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   277
     apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   278
    apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   279
   apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   280
  apply (case_tac "(OK x) +_fe (OK y)")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   281
   apply assumption
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   282
  apply (rename_tac z)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   283
  apply (subgoal_tac "OK z: err A")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   284
  apply (drule eq_order_le)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   285
    apply (erule Semilat.orderI [OF Semilat.intro])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   286
   apply (blast dest: plus_le_conv3) 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   287
  apply (erule subst)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   288
  apply (blast intro: Semilat.closedI [OF Semilat.intro] closedD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   289
  done 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   290
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   291
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   292
section {* semilat (err(Union AS)) *}
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   293
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   294
(* FIXME? *)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   295
lemma all_bex_swap_lemma [iff]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   296
  "(!x. (? y:A. x = f y) \<longrightarrow> P x) = (!y:A. P(f y))"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   297
  by blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   298
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   299
lemma closed_err_Union_lift2I: 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   300
  "\<lbrakk> !A:AS. closed (err A) (lift2 f); AS ~= {}; 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   301
      !A:AS.!B:AS. A~=B \<longrightarrow> (!a:A.!b:B. a +_f b = Err) \<rbrakk> 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   302
  \<Longrightarrow> closed (err(Union AS)) (lift2 f)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   303
apply (unfold closed_def err_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   304
apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   305
apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   306
apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   307
apply fast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   308
done 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   309
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   310
text {* 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   311
  If @{term "AS = {}"} the thm collapses to
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   312
  @{prop "order r & closed {Err} f & Err +_f Err = Err"}
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   313
  which may not hold 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   314
*}
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   315
lemma err_semilat_UnionI:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   316
  "\<lbrakk> !A:AS. err_semilat(A, r, f); AS ~= {}; 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   317
      !A:AS.!B:AS. A~=B \<longrightarrow> (!a:A.!b:B. ~ a <=_r b & a +_f b = Err) \<rbrakk> 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   318
  \<Longrightarrow> err_semilat(Union AS, r, f)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   319
apply (unfold semilat_def sl_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   320
apply (simp add: closed_err_Union_lift2I)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   321
apply (rule conjI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   322
 apply blast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   323
apply (simp add: err_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   324
apply (rule conjI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   325
 apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   326
 apply (rename_tac A a u B b)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   327
 apply (case_tac "A = B")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   328
  apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   329
 apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   330
apply (rule conjI)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   331
 apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   332
 apply (rename_tac A a u B b)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   333
 apply (case_tac "A = B")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   334
  apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   335
 apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   336
apply clarify
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   337
apply (rename_tac A ya yb B yd z C c a b)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   338
apply (case_tac "A = B")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   339
 apply (case_tac "A = C")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   340
  apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   341
 apply (rotate_tac -1)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   342
 apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   343
apply (rotate_tac -1)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   344
apply (case_tac "B = C")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   345
 apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   346
apply (rotate_tac -1)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   347
apply simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   348
done 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   349
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   350
end