src/HOL/MicroJava/BV/Typing_Framework_err.thy
author nipkow
Wed, 28 Mar 2001 13:40:06 +0200
changeset 11229 f417841385b7
child 11299 1d3d110c456e
permissions -rw-r--r--
Got rid of is_dfa
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11229
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/BV/Typing_Framework_err.thy
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     2
    ID:         $Id$
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     3
    Author:     Gerwin Klein
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     4
    Copyright   2000 TUM
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     5
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     6
*)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     7
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     8
header "Static and Dynamic Welltyping"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     9
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    10
theory Typing_Framework_err = Typing_Framework:
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    11
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    12
constdefs
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    13
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    14
dynamic_wt :: "'s ord => (nat => 's err => 's err) => (nat => nat list) => 
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    15
               's err list => bool"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    16
"dynamic_wt r step succs ts == welltyping (Err.le r) Err step succs ts"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    17
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    18
static_wt :: "'s ord => (nat => 's => bool) => 
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    19
              (nat => 's => 's) => (nat => nat list) =>  's list => bool"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    20
"static_wt r app step succs ts == 
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    21
  !p < size ts. app p (ts!p) & (!q:set(succs p). step p (ts!p) <=_r ts!q)"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    22
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    23
err_step :: "(nat => 's => bool) => (nat => 's => 's) => 
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    24
             (nat => 's err => 's err)"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    25
"err_step app step p == lift (%t. if app p t then OK (step p t) else Err)"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    26
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    27
non_empty :: "(nat => nat list) => bool"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    28
"non_empty succs == !p. succs p ~= []"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    29
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    30
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    31
lemma non_emptyD:
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    32
  "non_empty succs ==> ? q. q:set(succs p)"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    33
proof (unfold non_empty_def)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    34
  assume "!p. succs p ~= []"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    35
  hence "succs p ~= []" ..
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    36
  then obtain x xs where "succs p = x#xs"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    37
    by (auto simp add: neq_Nil_conv)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    38
  thus ?thesis 
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    39
    by auto
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    40
qed
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    41
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    42
lemma dynamic_imp_static:
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    43
  "[| bounded succs (size ts); non_empty succs;
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    44
      dynamic_wt r (err_step app step) succs ts |] 
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    45
  ==> static_wt r app step succs (map ok_val ts)"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    46
proof (unfold static_wt_def, intro strip, rule conjI)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    47
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    48
  assume b:  "bounded succs (size ts)"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    49
  assume n:  "non_empty succs"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    50
  assume wt: "dynamic_wt r (err_step app step) succs ts"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    51
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    52
  fix p 
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    53
  assume "p < length (map ok_val ts)"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    54
  hence lp: "p < length ts" by simp
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    55
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    56
  from wt lp
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    57
  have [intro?]: "!!p. p < length ts ==> ts ! p ~= Err" 
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    58
    by (unfold dynamic_wt_def welltyping_def) simp
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    59
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    60
  show app: "app p (map ok_val ts ! p)"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    61
  proof -
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    62
    from n
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    63
    obtain q where q: "q:set(succs p)"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    64
      by (auto dest: non_emptyD)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    65
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    66
    from wt lp q
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    67
    obtain s where
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    68
      OKp:  "ts ! p = OK s" and
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    69
      less: "err_step app step p (ts ! p) <=_(Err.le r) ts ! q"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    70
      by (unfold dynamic_wt_def welltyping_def stable_def) 
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    71
         (auto iff: not_Err_eq)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    72
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    73
    from lp b q
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    74
    have lq: "q < length ts"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    75
      by (unfold bounded_def) blast
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    76
    hence "ts!q ~= Err" ..
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    77
    then
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    78
    obtain s' where OKq: "ts ! q = OK s'"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    79
      by (auto iff: not_Err_eq)      
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    80
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    81
    with OKp less
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    82
    have "app p s"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    83
      by (simp add: err_step_def lift_def split: err.split_asm split_if_asm)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    84
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    85
    with lp OKp
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    86
    show ?thesis
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    87
      by simp
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    88
  qed
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    89
  
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    90
  show "!q:set(succs p). step p (map ok_val ts ! p) <=_r map ok_val ts ! q"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    91
  proof (intro strip)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    92
    fix q
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    93
    assume q: "q:set (succs p)"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    94
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    95
    from wt lp q
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    96
    obtain s where
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    97
      OKp:  "ts ! p = OK s" and
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    98
      less: "err_step app step p (ts ! p) <=_(Err.le r) ts ! q"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    99
      by (unfold dynamic_wt_def welltyping_def stable_def) 
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   100
         (auto iff: not_Err_eq)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   101
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   102
    from lp b q
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   103
    have lq: "q < length ts"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   104
      by (unfold bounded_def) blast
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   105
    hence "ts!q ~= Err" ..
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   106
    then
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   107
    obtain s' where OKq: "ts ! q = OK s'"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   108
      by (auto iff: not_Err_eq)      
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   109
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   110
    from lp lq OKp OKq app less
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   111
    show "step p (map ok_val ts ! p) <=_r map ok_val ts ! q"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   112
      by (simp add: err_step_def lift_def)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   113
  qed
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   114
qed
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   115
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   116
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   117
lemma static_imp_dynamic:
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   118
  "[| static_wt r app step succs ts; bounded succs (size ts) |] 
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   119
  ==> dynamic_wt r (err_step app step) succs (map OK ts)"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   120
proof (unfold dynamic_wt_def welltyping_def, intro strip, rule conjI)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   121
  assume bounded: "bounded succs (size ts)"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   122
  assume static:  "static_wt r app step succs ts"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   123
  fix p assume "p < length (map OK ts)" 
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   124
  hence p: "p < length ts" by simp
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   125
  thus "map OK ts ! p \<noteq> Err" by simp
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   126
  { fix q
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   127
    assume q: "q : set (succs p)"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   128
    with p static obtain 
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   129
      "app p (ts ! p)" "step p (ts ! p) <=_r ts ! q"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   130
      by (unfold static_wt_def) blast
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   131
    moreover
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   132
    from q p bounded
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   133
    have "q < size ts" by (unfold bounded_def) blast
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   134
    hence "map OK ts ! q = OK (ts!q)" by simp
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   135
    moreover
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   136
    have "p < size ts" by (rule p)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   137
    ultimately     
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   138
    have "err_step app step p (map OK ts ! p) <=_(Err.le r) map OK ts ! q"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   139
      by (simp add: err_step_def lift_def)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   140
  }
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   141
  thus "stable (Err.le r) (err_step app step) succs (map OK ts) p"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   142
    by (unfold stable_def) blast
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   143
qed
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   144
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
   145
end