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