| 
12516
 | 
     1  | 
(*  Title:      HOL/MicroJava/BV/Opt.thy
  | 
| 
10496
 | 
     2  | 
    ID:         $Id$
  | 
| 
 | 
     3  | 
    Author:     Tobias Nipkow
  | 
| 
 | 
     4  | 
    Copyright   2000 TUM
  | 
| 
 | 
     5  | 
  | 
| 
 | 
     6  | 
More about options
  | 
| 
 | 
     7  | 
*)
  | 
| 
 | 
     8  | 
  | 
| 
12911
 | 
     9  | 
header {* \isaheader{More about Options} *}
 | 
| 
10496
 | 
    10  | 
  | 
| 
16417
 | 
    11  | 
theory Opt imports Err begin
  | 
| 
10496
 | 
    12  | 
  | 
| 
 | 
    13  | 
constdefs
  | 
| 
13006
 | 
    14  | 
 le :: "'a ord \<Rightarrow> 'a option ord"
  | 
| 
 | 
    15  | 
"le r o1 o2 == case o2 of None \<Rightarrow> o1=None |
  | 
| 
 | 
    16  | 
                              Some y \<Rightarrow> (case o1 of None \<Rightarrow> True
  | 
| 
 | 
    17  | 
                                                  | Some x \<Rightarrow> x <=_r y)"
  | 
| 
10496
 | 
    18  | 
  | 
| 
13006
 | 
    19  | 
 opt :: "'a set \<Rightarrow> 'a option set"
  | 
| 
10496
 | 
    20  | 
"opt A == insert None {x . ? y:A. x = Some y}"
 | 
| 
 | 
    21  | 
  | 
| 
13006
 | 
    22  | 
 sup :: "'a ebinop \<Rightarrow> 'a option ebinop"
  | 
| 
10496
 | 
    23  | 
"sup f o1 o2 ==  
  | 
| 
13006
 | 
    24  | 
 case o1 of None \<Rightarrow> OK o2 | Some x \<Rightarrow> (case o2 of None \<Rightarrow> OK o1
  | 
| 
 | 
    25  | 
     | Some y \<Rightarrow> (case f x y of Err \<Rightarrow> Err | OK z \<Rightarrow> OK (Some z)))"
  | 
| 
10496
 | 
    26  | 
  | 
| 
13006
 | 
    27  | 
 esl :: "'a esl \<Rightarrow> 'a option esl"
  | 
| 
10496
 | 
    28  | 
"esl == %(A,r,f). (opt A, le r, sup f)"
  | 
| 
 | 
    29  | 
  | 
| 
 | 
    30  | 
lemma unfold_le_opt:
  | 
| 
 | 
    31  | 
  "o1 <=_(le r) o2 = 
  | 
| 
13006
 | 
    32  | 
  (case o2 of None \<Rightarrow> o1=None | 
  | 
| 
 | 
    33  | 
              Some y \<Rightarrow> (case o1 of None \<Rightarrow> True | Some x \<Rightarrow> x <=_r y))"
  | 
| 
10496
 | 
    34  | 
apply (unfold lesub_def le_def)
  | 
| 
 | 
    35  | 
apply (rule refl)
  | 
| 
 | 
    36  | 
done
  | 
| 
 | 
    37  | 
  | 
| 
 | 
    38  | 
lemma le_opt_refl:
  | 
| 
13006
 | 
    39  | 
  "order r \<Longrightarrow> o1 <=_(le r) o1"
  | 
| 
10496
 | 
    40  | 
by (simp add: unfold_le_opt split: option.split)
  | 
| 
 | 
    41  | 
  | 
| 
 | 
    42  | 
lemma le_opt_trans [rule_format]:
  | 
| 
13006
 | 
    43  | 
  "order r \<Longrightarrow> 
  | 
| 
 | 
    44  | 
   o1 <=_(le r) o2 \<longrightarrow> o2 <=_(le r) o3 \<longrightarrow> o1 <=_(le r) o3"
  | 
| 
10496
 | 
    45  | 
apply (simp add: unfold_le_opt split: option.split)
  | 
| 
 | 
    46  | 
apply (blast intro: order_trans)
  | 
| 
 | 
    47  | 
done
  | 
| 
 | 
    48  | 
  | 
| 
 | 
    49  | 
lemma le_opt_antisym [rule_format]:
  | 
| 
13006
 | 
    50  | 
  "order r \<Longrightarrow> o1 <=_(le r) o2 \<longrightarrow> o2 <=_(le r) o1 \<longrightarrow> o1=o2"
  | 
| 
10496
 | 
    51  | 
apply (simp add: unfold_le_opt split: option.split)
  | 
| 
 | 
    52  | 
apply (blast intro: order_antisym)
  | 
| 
 | 
    53  | 
done
  | 
| 
 | 
    54  | 
  | 
| 
 | 
    55  | 
lemma order_le_opt [intro!,simp]:
  | 
| 
13006
 | 
    56  | 
  "order r \<Longrightarrow> order(le r)"
  | 
| 
22068
 | 
    57  | 
apply (subst Semilat.order_def)
  | 
| 
10496
 | 
    58  | 
apply (blast intro: le_opt_refl le_opt_trans le_opt_antisym)
  | 
| 
 | 
    59  | 
done 
  | 
| 
 | 
    60  | 
  | 
| 
 | 
    61  | 
lemma None_bot [iff]: 
  | 
| 
 | 
    62  | 
  "None <=_(le r) ox"
  | 
| 
 | 
    63  | 
apply (unfold lesub_def le_def)
  | 
| 
 | 
    64  | 
apply (simp split: option.split)
  | 
| 
 | 
    65  | 
done 
  | 
| 
 | 
    66  | 
  | 
| 
 | 
    67  | 
lemma Some_le [iff]:
  | 
| 
 | 
    68  | 
  "(Some x <=_(le r) ox) = (? y. ox = Some y & x <=_r y)"
  | 
| 
 | 
    69  | 
apply (unfold lesub_def le_def)
  | 
| 
 | 
    70  | 
apply (simp split: option.split)
  | 
| 
 | 
    71  | 
done 
  | 
| 
 | 
    72  | 
  | 
| 
 | 
    73  | 
lemma le_None [iff]:
  | 
| 
 | 
    74  | 
  "(ox <=_(le r) None) = (ox = None)";
  | 
| 
 | 
    75  | 
apply (unfold lesub_def le_def)
  | 
| 
 | 
    76  | 
apply (simp split: option.split)
  | 
| 
 | 
    77  | 
done 
  | 
| 
 | 
    78  | 
  | 
| 
 | 
    79  | 
  | 
| 
 | 
    80  | 
lemma OK_None_bot [iff]:
  | 
| 
 | 
    81  | 
  "OK None <=_(Err.le (le r)) x"
  | 
| 
 | 
    82  | 
  by (simp add: lesub_def Err.le_def le_def split: option.split err.split)
  | 
| 
 | 
    83  | 
  | 
| 
 | 
    84  | 
lemma sup_None1 [iff]:
  | 
| 
 | 
    85  | 
  "x +_(sup f) None = OK x"
  | 
| 
 | 
    86  | 
  by (simp add: plussub_def sup_def split: option.split)
  | 
| 
 | 
    87  | 
  | 
| 
 | 
    88  | 
lemma sup_None2 [iff]:
  | 
| 
 | 
    89  | 
  "None +_(sup f) x = OK x"
  | 
| 
 | 
    90  | 
  by (simp add: plussub_def sup_def split: option.split)
  | 
| 
 | 
    91  | 
  | 
| 
 | 
    92  | 
  | 
| 
 | 
    93  | 
lemma None_in_opt [iff]:
  | 
| 
 | 
    94  | 
  "None : opt A"
  | 
| 
 | 
    95  | 
by (simp add: opt_def)
  | 
| 
 | 
    96  | 
  | 
| 
 | 
    97  | 
lemma Some_in_opt [iff]:
  | 
| 
 | 
    98  | 
  "(Some x : opt A) = (x:A)"
  | 
| 
 | 
    99  | 
apply (unfold opt_def)
  | 
| 
 | 
   100  | 
apply auto
  | 
| 
 | 
   101  | 
done 
  | 
| 
 | 
   102  | 
  | 
| 
 | 
   103  | 
  | 
| 
13062
 | 
   104  | 
lemma semilat_opt [intro, simp]:
  | 
| 
13006
 | 
   105  | 
  "\<And>L. err_semilat L \<Longrightarrow> err_semilat (Opt.esl L)"
  | 
| 
10496
 | 
   106  | 
proof (unfold Opt.esl_def Err.sl_def, simp add: split_tupled_all)
  | 
| 
 | 
   107  | 
  
  | 
| 
 | 
   108  | 
  fix A r f
  | 
| 
 | 
   109  | 
  assume s: "semilat (err A, Err.le r, lift2 f)"
  | 
| 
 | 
   110  | 
 
  | 
| 
 | 
   111  | 
  let ?A0 = "err A"
  | 
| 
 | 
   112  | 
  let ?r0 = "Err.le r"
  | 
| 
 | 
   113  | 
  let ?f0 = "lift2 f"
  | 
| 
 | 
   114  | 
  | 
| 
 | 
   115  | 
  from s
  | 
| 
 | 
   116  | 
  obtain
  | 
| 
 | 
   117  | 
    ord: "order ?r0" and
  | 
| 
 | 
   118  | 
    clo: "closed ?A0 ?f0" and
  | 
| 
 | 
   119  | 
    ub1: "\<forall>x\<in>?A0. \<forall>y\<in>?A0. x <=_?r0 x +_?f0 y" and
  | 
| 
 | 
   120  | 
    ub2: "\<forall>x\<in>?A0. \<forall>y\<in>?A0. y <=_?r0 x +_?f0 y" and
  | 
| 
 | 
   121  | 
    lub: "\<forall>x\<in>?A0. \<forall>y\<in>?A0. \<forall>z\<in>?A0. x <=_?r0 z \<and> y <=_?r0 z \<longrightarrow> x +_?f0 y <=_?r0 z"
  | 
| 
 | 
   122  | 
    by (unfold semilat_def) simp
  | 
| 
 | 
   123  | 
  | 
| 
 | 
   124  | 
  let ?A = "err (opt A)" 
  | 
| 
 | 
   125  | 
  let ?r = "Err.le (Opt.le r)"
  | 
| 
 | 
   126  | 
  let ?f = "lift2 (Opt.sup f)"
  | 
| 
 | 
   127  | 
  | 
| 
 | 
   128  | 
  from ord
  | 
| 
 | 
   129  | 
  have "order ?r"
  | 
| 
 | 
   130  | 
    by simp
  | 
| 
 | 
   131  | 
  | 
| 
 | 
   132  | 
  moreover
  | 
| 
 | 
   133  | 
  | 
| 
 | 
   134  | 
  have "closed ?A ?f"
  | 
| 
 | 
   135  | 
  proof (unfold closed_def, intro strip)
  | 
| 
 | 
   136  | 
    fix x y    
  | 
| 
 | 
   137  | 
    assume x: "x : ?A" 
  | 
| 
 | 
   138  | 
    assume y: "y : ?A" 
  | 
| 
 | 
   139  | 
  | 
| 
11085
 | 
   140  | 
    { fix a b
 | 
| 
10496
 | 
   141  | 
      assume ab: "x = OK a" "y = OK b"
  | 
| 
 | 
   142  | 
      
  | 
| 
 | 
   143  | 
      with x 
  | 
| 
13006
 | 
   144  | 
      have a: "\<And>c. a = Some c \<Longrightarrow> c : A"
  | 
| 
10496
 | 
   145  | 
        by (clarsimp simp add: opt_def)
  | 
| 
 | 
   146  | 
  | 
| 
 | 
   147  | 
      from ab y
  | 
| 
13006
 | 
   148  | 
      have b: "\<And>d. b = Some d \<Longrightarrow> d : A"
  | 
| 
10496
 | 
   149  | 
        by (clarsimp simp add: opt_def)
  | 
| 
 | 
   150  | 
      
  | 
| 
 | 
   151  | 
      { fix c d assume "a = Some c" "b = Some d"
 | 
| 
 | 
   152  | 
        with ab x y
  | 
| 
 | 
   153  | 
        have "c:A & d:A"
  | 
| 
 | 
   154  | 
          by (simp add: err_def opt_def Bex_def)
  | 
| 
 | 
   155  | 
        with clo
  | 
| 
 | 
   156  | 
        have "f c d : err A"
  | 
| 
 | 
   157  | 
          by (simp add: closed_def plussub_def err_def lift2_def)
  | 
| 
 | 
   158  | 
        moreover
  | 
| 
 | 
   159  | 
        fix z assume "f c d = OK z"
  | 
| 
 | 
   160  | 
        ultimately
  | 
| 
 | 
   161  | 
        have "z : A" by simp
  | 
| 
 | 
   162  | 
      } note f_closed = this    
  | 
| 
 | 
   163  | 
  | 
| 
 | 
   164  | 
      have "sup f a b : ?A"
  | 
| 
 | 
   165  | 
      proof (cases a)
  | 
| 
 | 
   166  | 
        case None
  | 
| 
 | 
   167  | 
        thus ?thesis
  | 
| 
 | 
   168  | 
          by (simp add: sup_def opt_def) (cases b, simp, simp add: b Bex_def)
  | 
| 
 | 
   169  | 
      next
  | 
| 
 | 
   170  | 
        case Some
  | 
| 
 | 
   171  | 
        thus ?thesis
  | 
| 
 | 
   172  | 
          by (auto simp add: sup_def opt_def Bex_def a b f_closed split: err.split option.split)
  | 
| 
 | 
   173  | 
      qed
  | 
| 
 | 
   174  | 
    }
  | 
| 
 | 
   175  | 
  | 
| 
 | 
   176  | 
    thus "x +_?f y : ?A"
  | 
| 
 | 
   177  | 
      by (simp add: plussub_def lift2_def split: err.split)
  | 
| 
 | 
   178  | 
  qed
  | 
| 
 | 
   179  | 
    
  | 
| 
 | 
   180  | 
  moreover
  | 
| 
 | 
   181  | 
  | 
| 
11085
 | 
   182  | 
  { fix a b c 
 | 
| 
 | 
   183  | 
    assume "a \<in> opt A" "b \<in> opt A" "a +_(sup f) b = OK c" 
  | 
| 
 | 
   184  | 
    moreover
  | 
| 
 | 
   185  | 
    from ord have "order r" by simp
  | 
| 
 | 
   186  | 
    moreover
  | 
| 
 | 
   187  | 
    { fix x y z
 | 
| 
 | 
   188  | 
      assume "x \<in> A" "y \<in> A" 
  | 
| 
 | 
   189  | 
      hence "OK x \<in> err A \<and> OK y \<in> err A" by simp
  | 
| 
 | 
   190  | 
      with ub1 ub2
  | 
| 
 | 
   191  | 
      have "(OK x) <=_(Err.le r) (OK x) +_(lift2 f) (OK y) \<and>
  | 
| 
 | 
   192  | 
            (OK y) <=_(Err.le r) (OK x) +_(lift2 f) (OK y)"
  | 
| 
 | 
   193  | 
        by blast
  | 
| 
 | 
   194  | 
      moreover
  | 
| 
 | 
   195  | 
      assume "x +_f y = OK z"
  | 
| 
 | 
   196  | 
      ultimately
  | 
| 
 | 
   197  | 
      have "x <=_r z \<and> y <=_r z"
  | 
| 
 | 
   198  | 
        by (auto simp add: plussub_def lift2_def Err.le_def lesub_def)
  | 
| 
 | 
   199  | 
    }
  | 
| 
 | 
   200  | 
    ultimately
  | 
| 
 | 
   201  | 
    have "a <=_(le r) c \<and> b <=_(le r) c"
  | 
| 
 | 
   202  | 
      by (auto simp add: sup_def le_def lesub_def plussub_def 
  | 
| 
 | 
   203  | 
               dest: order_refl split: option.splits err.splits)
  | 
| 
 | 
   204  | 
  }
  | 
| 
 | 
   205  | 
     
  | 
| 
 | 
   206  | 
  hence "(\<forall>x\<in>?A. \<forall>y\<in>?A. x <=_?r x +_?f y) \<and> (\<forall>x\<in>?A. \<forall>y\<in>?A. y <=_?r x +_?f y)"
  | 
| 
 | 
   207  | 
    by (auto simp add: lesub_def plussub_def Err.le_def lift2_def split: err.split)
  | 
| 
10496
 | 
   208  | 
  | 
| 
 | 
   209  | 
  moreover
  | 
| 
 | 
   210  | 
  | 
| 
 | 
   211  | 
  have "\<forall>x\<in>?A. \<forall>y\<in>?A. \<forall>z\<in>?A. x <=_?r z \<and> y <=_?r z \<longrightarrow> x +_?f y <=_?r z"
  | 
| 
 | 
   212  | 
  proof (intro strip, elim conjE)
  | 
| 
 | 
   213  | 
    fix x y z
  | 
| 
 | 
   214  | 
    assume xyz: "x : ?A" "y : ?A" "z : ?A"
  | 
| 
 | 
   215  | 
    assume xz: "x <=_?r z"
  | 
| 
 | 
   216  | 
    assume yz: "y <=_?r z"
  | 
| 
 | 
   217  | 
  | 
| 
 | 
   218  | 
    { fix a b c
 | 
| 
 | 
   219  | 
      assume ok: "x = OK a" "y = OK b" "z = OK c"
  | 
| 
 | 
   220  | 
  | 
| 
 | 
   221  | 
      { fix d e g
 | 
| 
 | 
   222  | 
        assume some: "a = Some d" "b = Some e" "c = Some g"
  | 
| 
 | 
   223  | 
        
  | 
| 
 | 
   224  | 
        with ok xyz
  | 
| 
 | 
   225  | 
        obtain "OK d:err A" "OK e:err A" "OK g:err A"
  | 
| 
 | 
   226  | 
          by simp
  | 
| 
 | 
   227  | 
        with lub
  | 
| 
13006
 | 
   228  | 
        have "\<lbrakk> (OK d) <=_(Err.le r) (OK g); (OK e) <=_(Err.le r) (OK g) \<rbrakk>
  | 
| 
 | 
   229  | 
          \<Longrightarrow> (OK d) +_(lift2 f) (OK e) <=_(Err.le r) (OK g)"
  | 
| 
10496
 | 
   230  | 
          by blast
  | 
| 
13006
 | 
   231  | 
        hence "\<lbrakk> d <=_r g; e <=_r g \<rbrakk> \<Longrightarrow> \<exists>y. d +_f e = OK y \<and> y <=_r g"
  | 
| 
10496
 | 
   232  | 
          by simp
  | 
| 
 | 
   233  | 
  | 
| 
 | 
   234  | 
        with ok some xyz xz yz
  | 
| 
 | 
   235  | 
        have "x +_?f y <=_?r z"
  | 
| 
 | 
   236  | 
          by (auto simp add: sup_def le_def lesub_def lift2_def plussub_def Err.le_def)
  | 
| 
 | 
   237  | 
      } note this [intro!]
  | 
| 
 | 
   238  | 
  | 
| 
 | 
   239  | 
      from ok xyz xz yz
  | 
| 
 | 
   240  | 
      have "x +_?f y <=_?r z"
  | 
| 
 | 
   241  | 
        by - (cases a, simp, cases b, simp, cases c, simp, blast)
  | 
| 
 | 
   242  | 
    }
  | 
| 
 | 
   243  | 
    
  | 
| 
 | 
   244  | 
    with xyz xz yz
  | 
| 
 | 
   245  | 
    show "x +_?f y <=_?r z"
  | 
| 
 | 
   246  | 
      by - (cases x, simp, cases y, simp, cases z, simp+)
  | 
| 
 | 
   247  | 
  qed
  | 
| 
 | 
   248  | 
  | 
| 
 | 
   249  | 
  ultimately
  | 
| 
 | 
   250  | 
  | 
| 
 | 
   251  | 
  show "semilat (?A,?r,?f)"
  | 
| 
 | 
   252  | 
    by (unfold semilat_def) simp
  | 
| 
 | 
   253  | 
qed 
  | 
| 
 | 
   254  | 
  | 
| 
 | 
   255  | 
lemma top_le_opt_Some [iff]: 
  | 
| 
 | 
   256  | 
  "top (le r) (Some T) = top r T"
  | 
| 
 | 
   257  | 
apply (unfold top_def)
  | 
| 
 | 
   258  | 
apply (rule iffI)
  | 
| 
 | 
   259  | 
 apply blast
  | 
| 
 | 
   260  | 
apply (rule allI)
  | 
| 
 | 
   261  | 
apply (case_tac "x")
  | 
| 
 | 
   262  | 
apply simp+
  | 
| 
 | 
   263  | 
done 
  | 
| 
 | 
   264  | 
  | 
| 
 | 
   265  | 
lemma Top_le_conv:
  | 
| 
13006
 | 
   266  | 
  "\<lbrakk> order r; top r T \<rbrakk> \<Longrightarrow> (T <=_r x) = (x = T)"
  | 
| 
10496
 | 
   267  | 
apply (unfold top_def)
  | 
| 
 | 
   268  | 
apply (blast intro: order_antisym)
  | 
| 
 | 
   269  | 
done 
  | 
| 
 | 
   270  | 
  | 
| 
 | 
   271  | 
  | 
| 
 | 
   272  | 
lemma acc_le_optI [intro!]:
  | 
| 
13006
 | 
   273  | 
  "acc r \<Longrightarrow> acc(le r)"
  | 
| 
10496
 | 
   274  | 
apply (unfold acc_def lesub_def le_def lesssub_def)
  | 
| 
22271
 | 
   275  | 
apply (simp add: wfP_eq_minimal split: option.split)
  | 
| 
10496
 | 
   276  | 
apply clarify
  | 
| 
 | 
   277  | 
apply (case_tac "? a. Some a : Q")
  | 
| 
 | 
   278  | 
 apply (erule_tac x = "{a . Some a : Q}" in allE)
 | 
| 
 | 
   279  | 
 apply blast
  | 
| 
 | 
   280  | 
apply (case_tac "x")
  | 
| 
 | 
   281  | 
 apply blast
  | 
| 
 | 
   282  | 
apply blast
  | 
| 
 | 
   283  | 
done 
  | 
| 
 | 
   284  | 
  | 
| 
 | 
   285  | 
lemma option_map_in_optionI:
  | 
| 
13006
 | 
   286  | 
  "\<lbrakk> ox : opt S; !x:S. ox = Some x \<longrightarrow> f x : S \<rbrakk> 
  | 
| 
 | 
   287  | 
  \<Longrightarrow> option_map f ox : opt S";
  | 
| 
10496
 | 
   288  | 
apply (unfold option_map_def)
  | 
| 
 | 
   289  | 
apply (simp split: option.split)
  | 
| 
 | 
   290  | 
apply blast
  | 
| 
 | 
   291  | 
done 
  | 
| 
 | 
   292  | 
  | 
| 
 | 
   293  | 
end
  |