src/HOL/MicroJava/BV/Opt.thy
author kleing
Tue, 12 Dec 2000 14:08:26 +0100
changeset 10650 114999ff8d19
parent 10496 f2d304bdf3cc
child 11085 b830bf10bf71
permissions -rw-r--r--
added direction dynamic ==> static
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     1
(*  Title:      HOL/BCV/Opt.thy
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     2
    ID:         $Id$
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     3
    Author:     Tobias Nipkow
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     4
    Copyright   2000 TUM
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     5
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     6
More about options
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     7
*)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     8
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
     9
header "More about Options"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    10
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    11
theory Opt = Err:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    12
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    13
constdefs
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    14
 le :: "'a ord => 'a option ord"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    15
"le r o1 o2 == case o2 of None => o1=None |
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    16
                              Some y => (case o1 of None => True |
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    17
                                                    Some x => x <=_r y)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    18
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    19
 opt :: "'a set => 'a option set"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    20
"opt A == insert None {x . ? y:A. x = Some y}"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    21
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    22
 sup :: "'a ebinop => 'a option ebinop"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    23
"sup f o1 o2 ==  
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    24
 case o1 of None => OK o2 | Some x => (case o2 of None => OK o1
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    25
                                              | Some y => (case f x y of Err => Err | OK z => OK (Some z)))"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    26
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    27
 esl :: "'a esl => 'a option esl"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    28
"esl == %(A,r,f). (opt A, le r, sup f)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    29
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    30
lemma unfold_le_opt:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    31
  "o1 <=_(le r) o2 = 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    32
  (case o2 of None => o1=None | 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    33
              Some y => (case o1 of None => True | Some x => x <=_r y))"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    34
apply (unfold lesub_def le_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    35
apply (rule refl)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    36
done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    37
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    38
lemma le_opt_refl:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    39
  "order r ==> o1 <=_(le r) o1"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    40
by (simp add: unfold_le_opt split: option.split)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    41
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    42
lemma le_opt_trans [rule_format]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    43
  "order r ==> 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    44
   o1 <=_(le r) o2 --> o2 <=_(le r) o3 --> o1 <=_(le r) o3"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    45
apply (simp add: unfold_le_opt split: option.split)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    46
apply (blast intro: order_trans)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    47
done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    48
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    49
lemma le_opt_antisym [rule_format]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    50
  "order r ==> o1 <=_(le r) o2 --> o2 <=_(le r) o1 --> o1=o2"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    51
apply (simp add: unfold_le_opt split: option.split)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    52
apply (blast intro: order_antisym)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    53
done
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    54
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    55
lemma order_le_opt [intro!,simp]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    56
  "order r ==> order(le r)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    57
apply (subst order_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    58
apply (blast intro: le_opt_refl le_opt_trans le_opt_antisym)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    59
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    60
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    61
lemma None_bot [iff]: 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    62
  "None <=_(le r) ox"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    63
apply (unfold lesub_def le_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    64
apply (simp split: option.split)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    65
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    66
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    67
lemma Some_le [iff]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    68
  "(Some x <=_(le r) ox) = (? y. ox = Some y & x <=_r y)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    69
apply (unfold lesub_def le_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    70
apply (simp split: option.split)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    71
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    72
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    73
lemma le_None [iff]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    74
  "(ox <=_(le r) None) = (ox = None)";
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    75
apply (unfold lesub_def le_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    76
apply (simp split: option.split)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    77
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    78
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    79
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    80
lemma OK_None_bot [iff]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    81
  "OK None <=_(Err.le (le r)) x"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    82
  by (simp add: lesub_def Err.le_def le_def split: option.split err.split)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    83
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    84
lemma sup_None1 [iff]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    85
  "x +_(sup f) None = OK x"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    86
  by (simp add: plussub_def sup_def split: option.split)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    87
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    88
lemma sup_None2 [iff]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    89
  "None +_(sup f) x = OK x"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    90
  by (simp add: plussub_def sup_def split: option.split)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    91
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    92
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    93
lemma None_in_opt [iff]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    94
  "None : opt A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    95
by (simp add: opt_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    96
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    97
lemma Some_in_opt [iff]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    98
  "(Some x : opt A) = (x:A)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
    99
apply (unfold opt_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   100
apply auto
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   101
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   102
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   103
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   104
lemma semilat_opt:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   105
  "!!L. err_semilat L ==> err_semilat (Opt.esl L)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   106
proof (unfold Opt.esl_def Err.sl_def, simp add: split_tupled_all)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   107
  
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   108
  fix A r f
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   109
  assume s: "semilat (err A, Err.le r, lift2 f)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   110
 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   111
  let ?A0 = "err A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   112
  let ?r0 = "Err.le r"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   113
  let ?f0 = "lift2 f"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   114
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   115
  from s
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   116
  obtain
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   117
    ord: "order ?r0" and
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   118
    clo: "closed ?A0 ?f0" and
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   119
    ub1: "\<forall>x\<in>?A0. \<forall>y\<in>?A0. x <=_?r0 x +_?f0 y" and
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   120
    ub2: "\<forall>x\<in>?A0. \<forall>y\<in>?A0. y <=_?r0 x +_?f0 y" and
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   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"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   122
    by (unfold semilat_def) simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   123
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   124
  let ?A = "err (opt A)" 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   125
  let ?r = "Err.le (Opt.le r)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   126
  let ?f = "lift2 (Opt.sup f)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   127
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   128
  from ord
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   129
  have "order ?r"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   130
    by simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   131
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   132
  moreover
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   133
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   134
  have "closed ?A ?f"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   135
  proof (unfold closed_def, intro strip)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   136
    fix x y    
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   137
    assume x: "x : ?A" 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   138
    assume y: "y : ?A" 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   139
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   140
    {
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   141
      fix a b
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   142
      assume ab: "x = OK a" "y = OK b"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   143
      
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   144
      with x 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   145
      have a: "!!c. a = Some c ==> c : A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   146
        by (clarsimp simp add: opt_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   147
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   148
      from ab y
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   149
      have b: "!!d. b = Some d ==> d : A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   150
        by (clarsimp simp add: opt_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   151
      
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   152
      { fix c d assume "a = Some c" "b = Some d"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   153
        with ab x y
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   154
        have "c:A & d:A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   155
          by (simp add: err_def opt_def Bex_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   156
        with clo
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   157
        have "f c d : err A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   158
          by (simp add: closed_def plussub_def err_def lift2_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   159
        moreover
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   160
        fix z assume "f c d = OK z"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   161
        ultimately
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   162
        have "z : A" by simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   163
      } note f_closed = this    
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   164
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   165
      have "sup f a b : ?A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   166
      proof (cases a)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   167
        case None
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   168
        thus ?thesis
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   169
          by (simp add: sup_def opt_def) (cases b, simp, simp add: b Bex_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   170
      next
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   171
        case Some
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   172
        thus ?thesis
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   173
          by (auto simp add: sup_def opt_def Bex_def a b f_closed split: err.split option.split)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   174
      qed
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   175
    }
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   176
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   177
    thus "x +_?f y : ?A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   178
      by (simp add: plussub_def lift2_def split: err.split)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   179
  qed
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   180
    
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   181
  moreover
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   182
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   183
  have "\<forall>x\<in>?A. \<forall>y\<in>?A. x <=_?r x +_?f y"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   184
  proof (intro strip)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   185
    fix x y
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   186
    assume x: "x : ?A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   187
    assume y: "y : ?A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   188
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   189
    show "x <=_?r x +_?f y"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   190
    proof -
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   191
      from ord 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   192
      have order_r: "order r" by simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   193
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   194
      { fix a b 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   195
        assume ok: "x = OK a" "y = OK b"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   196
        
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   197
        { fix c d 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   198
          assume some: "a = Some c" "b = Some d"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   199
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   200
          with x y ok
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   201
          obtain "c:A" "d:A" by simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   202
          then
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   203
          obtain "OK c : err A" "OK d : err A" by simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   204
          with ub1
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   205
          have OK: "OK c <=_(Err.le r) (OK c) +_(lift2 f) (OK d)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   206
            by blast        
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   207
          
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   208
          { fix e assume  "f c d = OK e"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   209
            with OK
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   210
            have "c <=_r e"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   211
              by (simp add: lesub_def Err.le_def plussub_def lift2_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   212
          }
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   213
          with ok some
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   214
          have ?thesis
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   215
            by (simp add: plussub_def sup_def lesub_def le_def lift2_def Err.le_def 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   216
                     split: err.split)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   217
        } note this [intro!]
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   218
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   219
        from ok
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   220
        have ?thesis
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   221
          by - (cases a, simp, cases b, simp add: order_r, blast)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   222
      }
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   223
      thus ?thesis
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   224
        by - (cases x, simp, cases y, simp+)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   225
    qed
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   226
  qed
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   227
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   228
  moreover
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   229
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   230
  have "\<forall>x\<in>?A. \<forall>y\<in>?A. y <=_?r x +_?f y" 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   231
  proof (intro strip)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   232
    fix x y
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   233
    assume x: "x : ?A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   234
    assume y: "y : ?A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   235
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   236
    show "y <=_?r x +_?f y"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   237
    proof -
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   238
      from ord 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   239
      have order_r: "order r" by simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   240
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   241
      { fix a b 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   242
        assume ok: "x = OK a" "y = OK b"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   243
        
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   244
        { fix c d 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   245
          assume some: "a = Some c" "b = Some d"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   246
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   247
          with x y ok
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   248
          obtain "c:A" "d:A" by simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   249
          then
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   250
          obtain "OK c : err A" "OK d : err A" by simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   251
          with ub2
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   252
          have OK: "OK d <=_(Err.le r) (OK c) +_(lift2 f) (OK d)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   253
            by blast        
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   254
          
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   255
          { fix e assume  "f c d = OK e"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   256
            with OK
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   257
            have "d <=_r e"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   258
              by (simp add: lesub_def Err.le_def plussub_def lift2_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   259
          }
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   260
          with ok some
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   261
          have ?thesis
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   262
            by (simp add: plussub_def sup_def lesub_def le_def lift2_def Err.le_def 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   263
                     split: err.split)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   264
        } note this [intro!]
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   265
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   266
        from ok
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   267
        have ?thesis
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   268
          by - (cases a, simp add: order_r, cases b, simp, blast)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   269
      }
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   270
      thus ?thesis
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   271
        by - (cases x, simp, cases y, simp+)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   272
    qed
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   273
  qed
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   274
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   275
  moreover
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   276
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   277
  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"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   278
  proof (intro strip, elim conjE)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   279
    fix x y z
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   280
    assume xyz: "x : ?A" "y : ?A" "z : ?A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   281
    assume xz: "x <=_?r z"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   282
    assume yz: "y <=_?r z"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   283
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   284
    { fix a b c
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   285
      assume ok: "x = OK a" "y = OK b" "z = OK c"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   286
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   287
      { fix d e g
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   288
        assume some: "a = Some d" "b = Some e" "c = Some g"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   289
        
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   290
        with ok xyz
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   291
        obtain "OK d:err A" "OK e:err A" "OK g:err A"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   292
          by simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   293
        with lub
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   294
        have "[| (OK d) <=_(Err.le r) (OK g); (OK e) <=_(Err.le r) (OK g) |]
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   295
          ==> (OK d) +_(lift2 f) (OK e) <=_(Err.le r) (OK g)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   296
          by blast
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   297
        hence "[| d <=_r g; e <=_r g |] ==> \<exists>y. d +_f e = OK y \<and> y <=_r g"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   298
          by simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   299
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   300
        with ok some xyz xz yz
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   301
        have "x +_?f y <=_?r z"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   302
          by (auto simp add: sup_def le_def lesub_def lift2_def plussub_def Err.le_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   303
      } note this [intro!]
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   304
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   305
      from ok xyz xz yz
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   306
      have "x +_?f y <=_?r z"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   307
        by - (cases a, simp, cases b, simp, cases c, simp, blast)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   308
    }
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   309
    
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   310
    with xyz xz yz
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   311
    show "x +_?f y <=_?r z"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   312
      by - (cases x, simp, cases y, simp, cases z, simp+)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   313
  qed
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   314
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   315
  ultimately
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   316
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   317
  show "semilat (?A,?r,?f)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   318
    by (unfold semilat_def) simp
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   319
qed 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   320
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   321
lemma top_le_opt_Some [iff]: 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   322
  "top (le r) (Some T) = top r T"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   323
apply (unfold top_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   324
apply (rule iffI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   325
 apply blast
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   326
apply (rule allI)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   327
apply (case_tac "x")
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   328
apply simp+
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   329
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   330
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   331
lemma Top_le_conv:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   332
  "[| order r; top r T |] ==> (T <=_r x) = (x = T)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   333
apply (unfold top_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   334
apply (blast intro: order_antisym)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   335
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   336
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   337
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   338
lemma acc_le_optI [intro!]:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   339
  "acc r ==> acc(le r)"
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   340
apply (unfold acc_def lesub_def le_def lesssub_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   341
apply (simp add: wf_eq_minimal split: option.split)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   342
apply clarify
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   343
apply (case_tac "? a. Some a : Q")
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   344
 apply (erule_tac x = "{a . Some a : Q}" in allE)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   345
 apply blast
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   346
apply (case_tac "x")
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   347
 apply blast
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   348
apply blast
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   349
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   350
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   351
lemma option_map_in_optionI:
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   352
  "[| ox : opt S; !x:S. ox = Some x --> f x : S |] 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   353
  ==> option_map f ox : opt S";
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   354
apply (unfold option_map_def)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   355
apply (simp split: option.split)
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   356
apply blast
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   357
done 
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   358
f2d304bdf3cc BCV integration (first step)
kleing
parents:
diff changeset
   359
end