src/HOL/MicroJava/BV/LBVComplete.thy
author kleing
Sun, 24 Mar 2002 14:05:53 +0100
changeset 13065 d6585b32412b
parent 13064 1f54a5fecaa6
child 13066 b57d926d1de2
permissions -rw-r--r--
more about match_exception_table
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8388
b66d3c49cb8d completeness of the lightweight bytecode verifier
kleing
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/BV/LBVComplete.thy
b66d3c49cb8d completeness of the lightweight bytecode verifier
kleing
parents:
diff changeset
     2
    ID:         $Id$
b66d3c49cb8d completeness of the lightweight bytecode verifier
kleing
parents:
diff changeset
     3
    Author:     Gerwin Klein
b66d3c49cb8d completeness of the lightweight bytecode verifier
kleing
parents:
diff changeset
     4
    Copyright   2000 Technische Universitaet Muenchen
9054
0e48e7d4d4f9 minor tuning for pdf documents
kleing
parents: 9012
diff changeset
     5
*)
8388
b66d3c49cb8d completeness of the lightweight bytecode verifier
kleing
parents:
diff changeset
     6
12911
704713ca07ea new document
kleing
parents: 12516
diff changeset
     7
header {* \isaheader{Completeness of the LBV} *}
8388
b66d3c49cb8d completeness of the lightweight bytecode verifier
kleing
parents:
diff changeset
     8
13064
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
     9
theory LBVComplete = LBVSpec + Typing_Framework:
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
    10
8388
b66d3c49cb8d completeness of the lightweight bytecode verifier
kleing
parents:
diff changeset
    11
constdefs
13064
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    12
  contains_targets :: "['s steptype, 's certificate, 's option list, nat, nat] \<Rightarrow> bool"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    13
  "contains_targets step cert phi pc n \<equiv>
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    14
  \<forall>(pc',s') \<in> set (step pc (OK (phi!pc))). pc' \<noteq> pc+1 \<and> pc' < n \<longrightarrow> cert!pc' = phi!pc'"
8388
b66d3c49cb8d completeness of the lightweight bytecode verifier
kleing
parents:
diff changeset
    15
13064
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    16
  fits :: "['s steptype, 's certificate, 's option list, nat] \<Rightarrow> bool"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    17
  "fits step cert phi n \<equiv> \<forall>pc. pc < n \<longrightarrow> 
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    18
                               contains_targets step cert phi pc n \<and>
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    19
                               (cert!pc = None \<or> cert!pc = phi!pc)"
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
    20
13064
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    21
  is_target :: "['s steptype, 's option list, nat, nat] \<Rightarrow> bool" 
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    22
  "is_target step phi pc' n \<equiv>
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    23
     \<exists>pc s'. pc' \<noteq> pc+1 \<and> pc < n \<and> (pc',s') \<in> set (step pc (OK (phi!pc)))"
8388
b66d3c49cb8d completeness of the lightweight bytecode verifier
kleing
parents:
diff changeset
    24
13064
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    25
  make_cert :: "['s steptype, 's option list, nat] \<Rightarrow> 's certificate"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    26
  "make_cert step phi n \<equiv> 
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    27
     map (\<lambda>pc. if is_target step phi pc n then phi!pc else None) [0..n(]"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    28
8388
b66d3c49cb8d completeness of the lightweight bytecode verifier
kleing
parents:
diff changeset
    29
  
9559
kleing
parents: 9549
diff changeset
    30
lemmas [simp del] = split_paired_Ex
kleing
parents: 9549
diff changeset
    31
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    32
lemma make_cert_target:
13064
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    33
  "\<lbrakk> pc < n; is_target step phi pc n \<rbrakk> \<Longrightarrow> make_cert step phi n ! pc = phi!pc"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    34
  by (simp add: make_cert_def)
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
    35
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    36
lemma make_cert_approx:
13064
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    37
  "\<lbrakk> pc < n; make_cert step phi n ! pc \<noteq> phi!pc \<rbrakk> \<Longrightarrow> 
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    38
   make_cert step phi n ! pc = None"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    39
  by (auto simp add: make_cert_def)
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
    40
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    41
lemma make_cert_contains_targets:
13064
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    42
  "pc < n \<Longrightarrow> contains_targets step (make_cert step phi n) phi pc n"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    43
proof (unfold contains_targets_def, clarify)
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    44
  fix pc' s'
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    45
  assume "pc < n"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    46
         "(pc',s') \<in> set (step pc (OK (phi ! pc)))"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    47
         "pc' \<noteq> pc+1" and
13064
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    48
    pc': "pc' < n"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    49
  hence "is_target step phi pc' n"  by (auto simp add: is_target_def)
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    50
  with pc' show "make_cert step phi n ! pc' = phi ! pc'" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    51
    by (auto intro: make_cert_target)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
    52
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
    53
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    54
theorem fits_make_cert:
13064
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    55
  "fits step (make_cert step phi n) phi n"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    56
  by (auto dest: make_cert_approx simp add: fits_def make_cert_contains_targets)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
    57
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
    58
lemma fitsD: 
13064
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    59
  "\<lbrakk> fits step cert phi n; (pc',s') \<in> set (step pc (OK (phi ! pc)));
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    60
      pc' \<noteq> Suc pc; pc < n; pc' < n \<rbrakk>
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
    61
  \<Longrightarrow> cert!pc' = phi!pc'"
13064
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    62
  by (auto simp add: fits_def contains_targets_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
    63
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
    64
lemma fitsD2:
13064
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    65
   "\<lbrakk> fits step cert phi n; pc < n; cert!pc = Some s \<rbrakk>
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
    66
  \<Longrightarrow> cert!pc = phi!pc"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    67
  by (auto simp add: fits_def)
13064
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    68
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    69
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    70
lemma merge_mono:
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    71
  assumes merge: "merge cert f r pc ss1 x = OK s1"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    72
  assumes less:  "ss2 <=|Err.le (Opt.le r)| ss1"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    73
  shows "\<exists>s2. merge cert f r pc ss2 x = s2 \<and> s2 \<le>|r (OK s1)"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    74
proof-
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    75
  from merge
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    76
  obtain "\<forall>(pc',s')\<in>set ss1. pc' \<noteq> pc+1 \<longrightarrow> s' \<le>|r (OK (cert!pc'))" and
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    77
         "(map snd [(p',t')\<in>ss1 . p' = pc+1] ++|f x) = OK s1"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    78
    sorry
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    79
  show ?thesis sorry
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    80
qed
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    81
  
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    82
  
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    83
lemma stable_wtl:
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    84
  assumes stable: "stable (Err.le (Opt.le r)) step (map OK phi) pc"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    85
  assumes fits:   "fits step cert phi n"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    86
  assumes pc:     "pc < length phi"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    87
  shows "wtl_inst cert f r step pc (phi!pc) \<noteq> Err"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    88
proof -
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    89
  from pc have [simp]: "map OK phi ! pc = OK (phi!pc)" by simp
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    90
  from stable 
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    91
  have "\<forall>(q,s')\<in>set (step pc (OK (phi!pc))). s' \<le>|r (map OK phi!q)"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    92
    by (simp add: stable_def)
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    93
  
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    94
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    95
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    96
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    97
lemma wtl_inst_mono:
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    98
  assumes wtl:  "wtl_inst cert f r step pc s1 = OK s1'"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
    99
  assumes fits: "fits step cert phi n"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
   100
  assumes pc:   "pc < n" 
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
   101
  assumes less: "OK s2 \<le>|r (OK s1)"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
   102
  shows "\<exists>s2'. wtl_inst cert f r step pc s2 = OK s2' \<and> OK s2' \<le>|r (OK s1')"
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
   103
apply (simp add: wtl_inst_def)
1f54a5fecaa6 first steps in semilattices..
kleing
parents: 13006
diff changeset
   104
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   105
                           
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   106
lemma wtl_inst_mono:
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   107
  "\<lbrakk> wtl_inst i G rT s1 cert mxs mpc pc = OK s1'; fits ins cert phi; 
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   108
      pc < length ins;  G \<turnstile> s2 <=' s1; i = ins!pc \<rbrakk> \<Longrightarrow>
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   109
  \<exists> s2'. wtl_inst (ins!pc) G rT s2 cert mxs mpc pc = OK s2' \<and> (G \<turnstile> s2' <=' s1')"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   110
proof -
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   111
  assume pc:   "pc < length ins" "i = ins!pc"
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   112
  assume wtl:  "wtl_inst i G rT s1 cert mxs mpc pc = OK s1'"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   113
  assume fits: "fits ins cert phi"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   114
  assume G:    "G \<turnstile> s2 <=' s1"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   115
  
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   116
  let "?eff s" = "eff i G s"
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   117
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   118
  from wtl G
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   119
  have app: "app i G mxs rT s2" by (auto simp add: wtl_inst_OK app_mono)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   120
  
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   121
  from wtl G
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   122
  have eff: "G \<turnstile> ?eff s2 <=' ?eff s1" 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   123
    by (auto intro: eff_mono simp add: wtl_inst_OK)
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   124
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
   125
  { also
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   126
    fix pc'
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
   127
    assume "pc' \<in> set (succs i pc)" "pc' \<noteq> pc+1"
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
   128
    with wtl 
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   129
    have "G \<turnstile> ?eff s1 <=' cert!pc'"
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
   130
      by (auto simp add: wtl_inst_OK) 
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   131
    finally
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   132
    have "G\<turnstile> ?eff s2 <=' cert!pc'" .
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   133
  } note cert = this
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   134
    
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   135
  have "\<exists>s2'. wtl_inst i G rT s2 cert mxs mpc pc = OK s2' \<and> G \<turnstile> s2' <=' s1'"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   136
  proof (cases "pc+1 \<in> set (succs i pc)")
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   137
    case True
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   138
    with wtl
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   139
    have s1': "s1' = ?eff s1" by (simp add: wtl_inst_OK)
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   140
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   141
    have "wtl_inst i G rT s2 cert mxs mpc pc = OK (?eff s2) \<and> G \<turnstile> ?eff s2 <=' s1'" 
9580
kleing
parents: 9559
diff changeset
   142
      (is "?wtl \<and> ?G")
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   143
    proof
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   144
      from True s1'
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   145
      show ?G by (auto intro: eff)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   146
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   147
      from True app wtl
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   148
      show ?wtl
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
   149
        by (clarsimp intro!: cert simp add: wtl_inst_OK)
9376
c32c5696ec2a flat instruction set
kleing
parents: 9260
diff changeset
   150
    qed
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   151
    thus ?thesis by blast
9376
c32c5696ec2a flat instruction set
kleing
parents: 9260
diff changeset
   152
  next
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   153
    case False
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   154
    with wtl
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
   155
    have "s1' = cert ! Suc pc" by (simp add: wtl_inst_OK)
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   156
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   157
    with False app wtl
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   158
    have "wtl_inst i G rT s2 cert mxs mpc pc = OK s1' \<and> G \<turnstile> s1' <=' s1'"
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
   159
      by (clarsimp intro!: cert simp add: wtl_inst_OK)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   160
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   161
    thus ?thesis by blast
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   162
  qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   163
  
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   164
  with pc show ?thesis by simp
9376
c32c5696ec2a flat instruction set
kleing
parents: 9260
diff changeset
   165
qed
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   166
9559
kleing
parents: 9549
diff changeset
   167
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   168
lemma wtl_cert_mono:
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   169
  "\<lbrakk> wtl_cert i G rT s1 cert mxs mpc pc = OK s1'; fits ins cert phi; 
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   170
      pc < length ins; G \<turnstile> s2 <=' s1; i = ins!pc \<rbrakk> \<Longrightarrow>
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   171
  \<exists> s2'. wtl_cert (ins!pc) G rT s2 cert mxs mpc pc = OK s2' \<and> (G \<turnstile> s2' <=' s1')"
9559
kleing
parents: 9549
diff changeset
   172
proof -
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   173
  assume wtl:  "wtl_cert i G rT s1 cert mxs mpc pc = OK s1'" and
9559
kleing
parents: 9549
diff changeset
   174
         fits: "fits ins cert phi" "pc < length ins"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   175
               "G \<turnstile> s2 <=' s1" "i = ins!pc"
9559
kleing
parents: 9549
diff changeset
   176
kleing
parents: 9549
diff changeset
   177
  show ?thesis
9664
4cae97480a6d open cases;
wenzelm
parents: 9594
diff changeset
   178
  proof (cases (open) "cert!pc")
9559
kleing
parents: 9549
diff changeset
   179
    case None
9580
kleing
parents: 9559
diff changeset
   180
    with wtl fits
kleing
parents: 9559
diff changeset
   181
    show ?thesis 
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   182
      by - (rule wtl_inst_mono [elim_format], (simp add: wtl_cert_def)+)
9559
kleing
parents: 9549
diff changeset
   183
  next
kleing
parents: 9549
diff changeset
   184
    case Some
9580
kleing
parents: 9559
diff changeset
   185
    with wtl fits
9559
kleing
parents: 9549
diff changeset
   186
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   187
    have G: "G \<turnstile> s2 <=' (Some a)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   188
      by - (rule sup_state_opt_trans, auto simp add: wtl_cert_def split: if_splits)
9559
kleing
parents: 9549
diff changeset
   189
kleing
parents: 9549
diff changeset
   190
    from Some wtl
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   191
    have "wtl_inst i G rT (Some a) cert mxs mpc pc = OK s1'" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   192
      by (simp add: wtl_cert_def split: if_splits)
9559
kleing
parents: 9549
diff changeset
   193
kleing
parents: 9549
diff changeset
   194
    with G fits
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   195
    have "\<exists> s2'. wtl_inst (ins!pc) G rT (Some a) cert mxs mpc pc = OK s2' \<and> 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   196
                 (G \<turnstile> s2' <=' s1')"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   197
      by - (rule wtl_inst_mono, (simp add: wtl_cert_def)+)
9559
kleing
parents: 9549
diff changeset
   198
    
9580
kleing
parents: 9559
diff changeset
   199
    with Some G
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   200
    show ?thesis by (simp add: wtl_cert_def)
9559
kleing
parents: 9549
diff changeset
   201
  qed
kleing
parents: 9549
diff changeset
   202
qed
kleing
parents: 9549
diff changeset
   203
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   204
 
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   205
lemma wt_instr_imp_wtl_inst:
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   206
  "\<lbrakk> wt_instr (ins!pc) G rT phi mxs max_pc pc; fits ins cert phi;
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   207
      pc < length ins; length ins = max_pc \<rbrakk> \<Longrightarrow> 
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   208
  wtl_inst (ins!pc) G rT (phi!pc) cert mxs max_pc pc \<noteq> Err"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   209
 proof -
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   210
  assume wt:   "wt_instr (ins!pc) G rT phi mxs max_pc pc" 
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   211
  assume fits: "fits ins cert phi"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   212
  assume pc:   "pc < length ins" "length ins = max_pc"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   213
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   214
  from wt 
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   215
  have app: "app (ins!pc) G mxs rT (phi!pc)" by (simp add: wt_instr_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   216
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   217
  from wt pc
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   218
  have pc': "\<And>pc'. pc' \<in> set (succs (ins!pc) pc) \<Longrightarrow> pc' < length ins"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   219
    by (simp add: wt_instr_def)
9376
c32c5696ec2a flat instruction set
kleing
parents: 9260
diff changeset
   220
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   221
  let ?s' = "eff (ins!pc) G (phi!pc)"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   222
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   223
  from wt fits pc
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   224
  have cert: "\<And>pc'. \<lbrakk>pc' \<in> set (succs (ins!pc) pc); pc' < max_pc; pc' \<noteq> pc+1\<rbrakk> 
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   225
    \<Longrightarrow> G \<turnstile> ?s' <=' cert!pc'"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   226
    by (auto dest: fitsD simp add: wt_instr_def)     
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   227
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   228
  from app pc cert pc'
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   229
  show ?thesis
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
   230
    by (auto simp add: wtl_inst_OK)
9376
c32c5696ec2a flat instruction set
kleing
parents: 9260
diff changeset
   231
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   232
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   233
lemma wt_less_wtl:
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   234
  "\<lbrakk> wt_instr (ins!pc) G rT phi mxs max_pc pc; 
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   235
      wtl_inst (ins!pc) G rT (phi!pc) cert mxs max_pc pc = OK s;
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   236
      fits ins cert phi; Suc pc < length ins; length ins = max_pc \<rbrakk> \<Longrightarrow> 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   237
  G \<turnstile> s <=' phi ! Suc pc"
9376
c32c5696ec2a flat instruction set
kleing
parents: 9260
diff changeset
   238
proof -
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   239
  assume wt:   "wt_instr (ins!pc) G rT phi mxs max_pc pc" 
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   240
  assume wtl:  "wtl_inst (ins!pc) G rT (phi!pc) cert mxs max_pc pc = OK s"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   241
  assume fits: "fits ins cert phi"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   242
  assume pc:   "Suc pc < length ins" "length ins = max_pc"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   243
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   244
  { assume suc: "Suc pc \<in> set (succs (ins ! pc) pc)"
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   245
    with wtl have "s = eff (ins!pc) G (phi!pc)"
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
   246
      by (simp add: wtl_inst_OK)
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   247
    also from suc wt have "G \<turnstile> \<dots> <=' phi!Suc pc"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   248
      by (simp add: wt_instr_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   249
    finally have ?thesis .
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   250
  }
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   251
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   252
  moreover
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   253
  { assume "Suc pc \<notin> set (succs (ins ! pc) pc)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   254
    
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   255
    with wtl
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   256
    have "s = cert!Suc pc" 
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
   257
      by (simp add: wtl_inst_OK)
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   258
    with fits pc
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   259
    have ?thesis
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   260
      by - (cases "cert!Suc pc", simp, drule fitsD2, assumption+, simp)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   261
  }
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   262
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   263
  ultimately
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   264
  show ?thesis by blast
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   265
qed
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   266
  
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   267
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   268
lemma wt_instr_imp_wtl_cert:
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   269
  "\<lbrakk> wt_instr (ins!pc) G rT phi mxs max_pc pc; fits ins cert phi;
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   270
      pc < length ins; length ins = max_pc \<rbrakk> \<Longrightarrow> 
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   271
  wtl_cert (ins!pc) G rT (phi!pc) cert mxs max_pc pc \<noteq> Err"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   272
proof -
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   273
  assume "wt_instr (ins!pc) G rT phi mxs max_pc pc" and 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   274
   fits: "fits ins cert phi" and
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   275
     pc: "pc < length ins" and
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   276
         "length ins = max_pc"
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   277
  
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   278
  hence wtl: "wtl_inst (ins!pc) G rT (phi!pc) cert mxs max_pc pc \<noteq> Err"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   279
    by (rule wt_instr_imp_wtl_inst)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   280
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   281
  hence "cert!pc = None \<Longrightarrow> ?thesis"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   282
    by (simp add: wtl_cert_def)
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   283
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   284
  moreover
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   285
  { fix s
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   286
    assume c: "cert!pc = Some s"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   287
    with fits pc
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   288
    have "cert!pc = phi!pc"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   289
      by (rule fitsD2)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   290
    from this c wtl
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   291
    have ?thesis
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   292
      by (clarsimp simp add: wtl_cert_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   293
  }
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   294
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   295
  ultimately
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   296
  show ?thesis by blast
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   297
qed
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   298
  
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   299
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   300
lemma wt_less_wtl_cert:
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   301
  "\<lbrakk> wt_instr (ins!pc) G rT phi mxs max_pc pc; 
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   302
      wtl_cert (ins!pc) G rT (phi!pc) cert mxs max_pc pc = OK s;
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   303
      fits ins cert phi; Suc pc < length ins; length ins = max_pc \<rbrakk> \<Longrightarrow> 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   304
  G \<turnstile> s <=' phi ! Suc pc"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   305
proof -
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   306
  assume wtl: "wtl_cert (ins!pc) G rT (phi!pc) cert mxs max_pc pc = OK s"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   307
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   308
  assume wti: "wt_instr (ins!pc) G rT phi mxs max_pc pc"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   309
              "fits ins cert phi" 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   310
              "Suc pc < length ins" "length ins = max_pc"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   311
  
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   312
  { assume "cert!pc = None"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   313
    with wtl
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   314
    have "wtl_inst (ins!pc) G rT (phi!pc) cert mxs max_pc pc = OK s"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   315
      by (simp add: wtl_cert_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   316
    with wti
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   317
    have ?thesis
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   318
      by - (rule wt_less_wtl)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   319
  }
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   320
  moreover
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   321
  { fix t
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   322
    assume c: "cert!pc = Some t"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   323
    with wti
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   324
    have "cert!pc = phi!pc"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   325
      by - (rule fitsD2, simp+)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   326
    with c wtl
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   327
    have "wtl_inst (ins!pc) G rT (phi!pc) cert mxs max_pc pc = OK s"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   328
      by (simp add: wtl_cert_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   329
    with wti
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   330
    have ?thesis
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   331
      by - (rule wt_less_wtl)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   332
  }   
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   333
    
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   334
  ultimately
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   335
  show ?thesis by blast
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   336
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   337
9559
kleing
parents: 9549
diff changeset
   338
text {*
kleing
parents: 9549
diff changeset
   339
  \medskip
kleing
parents: 9549
diff changeset
   340
  Main induction over the instruction list.
kleing
parents: 9549
diff changeset
   341
*}
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   342
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   343
theorem wt_imp_wtl_inst_list:
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   344
"\<forall> pc. (\<forall>pc'. pc' < length all_ins \<longrightarrow> 
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   345
        wt_instr (all_ins ! pc') G rT phi mxs (length all_ins) pc') \<longrightarrow>
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   346
       fits all_ins cert phi \<longrightarrow> 
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   347
       (\<exists>l. pc = length l \<and> all_ins = l@ins) \<longrightarrow>  
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   348
       pc < length all_ins \<longrightarrow>      
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   349
       (\<forall> s. (G \<turnstile> s <=' (phi!pc)) \<longrightarrow> 
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   350
             wtl_inst_list ins G rT cert mxs (length all_ins) pc s \<noteq> Err)" 
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   351
(is "\<forall>pc. ?wt \<longrightarrow> ?fits \<longrightarrow> ?l pc ins \<longrightarrow> ?len pc \<longrightarrow> ?wtl pc ins"  
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   352
 is "\<forall>pc. ?C pc ins" is "?P ins") 
9559
kleing
parents: 9549
diff changeset
   353
proof (induct "?P" "ins")
kleing
parents: 9549
diff changeset
   354
  case Nil
kleing
parents: 9549
diff changeset
   355
  show "?P []" by simp
kleing
parents: 9549
diff changeset
   356
next
kleing
parents: 9549
diff changeset
   357
  fix i ins'
kleing
parents: 9549
diff changeset
   358
  assume Cons: "?P ins'"
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   359
9559
kleing
parents: 9549
diff changeset
   360
  show "?P (i#ins')" 
kleing
parents: 9549
diff changeset
   361
  proof (intro allI impI, elim exE conjE)
kleing
parents: 9549
diff changeset
   362
    fix pc s l
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   363
    assume wt  : "\<forall>pc'. pc' < length all_ins \<longrightarrow> 
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   364
                        wt_instr (all_ins ! pc') G rT phi mxs (length all_ins) pc'"
9559
kleing
parents: 9549
diff changeset
   365
    assume fits: "fits all_ins cert phi"
kleing
parents: 9549
diff changeset
   366
    assume l   : "pc < length all_ins"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   367
    assume G   : "G \<turnstile> s <=' phi ! pc"
9559
kleing
parents: 9549
diff changeset
   368
    assume pc  : "all_ins = l@i#ins'" "pc = length l"
kleing
parents: 9549
diff changeset
   369
    hence  i   : "all_ins ! pc = i"
kleing
parents: 9549
diff changeset
   370
      by (simp add: nth_append)
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   371
9559
kleing
parents: 9549
diff changeset
   372
    from l wt
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   373
    have wti: "wt_instr (all_ins!pc) G rT phi mxs (length all_ins) pc" by blast
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   374
9559
kleing
parents: 9549
diff changeset
   375
    with fits l 
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   376
    have c: "wtl_cert (all_ins!pc) G rT (phi!pc) cert mxs (length all_ins) pc \<noteq> Err"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   377
      by - (drule wt_instr_imp_wtl_cert, auto)
9559
kleing
parents: 9549
diff changeset
   378
    
kleing
parents: 9549
diff changeset
   379
    from Cons
kleing
parents: 9549
diff changeset
   380
    have "?C (Suc pc) ins'" by blast
kleing
parents: 9549
diff changeset
   381
    with wt fits pc
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   382
    have IH: "?len (Suc pc) \<longrightarrow> ?wtl (Suc pc) ins'" by auto
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   383
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   384
    show "wtl_inst_list (i#ins') G rT cert mxs (length all_ins) pc s \<noteq> Err" 
9559
kleing
parents: 9549
diff changeset
   385
    proof (cases "?len (Suc pc)")
kleing
parents: 9549
diff changeset
   386
      case False
kleing
parents: 9549
diff changeset
   387
      with pc
kleing
parents: 9549
diff changeset
   388
      have "ins' = []" by simp
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   389
      with G i c fits l
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   390
      show ?thesis by (auto dest: wtl_cert_mono)
9559
kleing
parents: 9549
diff changeset
   391
    next
kleing
parents: 9549
diff changeset
   392
      case True
kleing
parents: 9549
diff changeset
   393
      with IH
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   394
      have wtl: "?wtl (Suc pc) ins'" by blast
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   395
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   396
      from c wti fits l True
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   397
      obtain s'' where
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   398
        "wtl_cert (all_ins!pc) G rT (phi!pc) cert mxs (length all_ins) pc = OK s''"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   399
        "G \<turnstile> s'' <=' phi ! Suc pc"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   400
        by clarsimp (drule wt_less_wtl_cert, auto)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   401
      moreover
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   402
      from calculation fits G l
9559
kleing
parents: 9549
diff changeset
   403
      obtain s' where
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   404
        c': "wtl_cert (all_ins!pc) G rT s cert mxs (length all_ins) pc = OK s'" and
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   405
        "G \<turnstile> s' <=' s''"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   406
        by - (drule wtl_cert_mono, auto)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   407
      ultimately
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   408
      have G': "G \<turnstile> s' <=' phi ! Suc pc" 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   409
        by - (rule sup_state_opt_trans)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   410
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   411
      with wtl
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   412
      have "wtl_inst_list ins' G rT cert mxs (length all_ins) (Suc pc) s' \<noteq> Err"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   413
        by auto
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   414
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   415
      with i c'
9559
kleing
parents: 9549
diff changeset
   416
      show ?thesis by auto
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   417
    qed
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   418
  qed
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   419
qed
9559
kleing
parents: 9549
diff changeset
   420
  
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   421
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   422
lemma fits_imp_wtl_method_complete:
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   423
  "\<lbrakk> wt_method G C pTs rT mxs mxl ins phi; fits ins cert phi \<rbrakk> 
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   424
  \<Longrightarrow> wtl_method G C pTs rT mxs mxl ins cert"
9594
42d11e0a7a8b Convert.thy now in Isar, tuned
kleing
parents: 9580
diff changeset
   425
by (simp add: wt_method_def wtl_method_def)
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   426
   (rule wt_imp_wtl_inst_list [rule_format, elim_format], auto simp add: wt_start_def) 
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   427
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   428
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   429
lemma wtl_method_complete:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   430
  "wt_method G C pTs rT mxs mxl ins phi 
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   431
  \<Longrightarrow> wtl_method G C pTs rT mxs mxl ins (make_cert ins phi)"
9580
kleing
parents: 9559
diff changeset
   432
proof -
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   433
  assume "wt_method G C pTs rT mxs mxl ins phi" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   434
  moreover
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   435
  have "fits ins (make_cert ins phi) phi"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   436
    by (rule fits_make_cert)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   437
  ultimately
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   438
  show ?thesis 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   439
    by (rule fits_imp_wtl_method_complete)
9580
kleing
parents: 9559
diff changeset
   440
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   441
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8388
diff changeset
   442
10628
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   443
theorem wtl_complete:
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   444
  "wt_jvm_prog G Phi \<Longrightarrow> wtl_jvm_prog G (make_Cert G Phi)"
10628
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   445
proof -
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   446
  assume wt: "wt_jvm_prog G Phi"
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   447
   
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   448
  { fix C S fs mdecls sig rT code
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   449
    assume "(C,S,fs,mdecls) \<in> set G" "(sig,rT,code) \<in> set mdecls"
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   450
    moreover
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   451
    from wt obtain wf_mb where "wf_prog wf_mb G" 
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   452
      by (blast dest: wt_jvm_progD)
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   453
    ultimately
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   454
    have "method (G,C) sig = Some (C,rT,code)"
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   455
      by (simp add: methd)
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   456
  } note this [simp]
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   457
 
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   458
  from wt
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   459
  show ?thesis
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   460
    apply (clarsimp simp add: wt_jvm_prog_def wtl_jvm_prog_def wf_prog_def wf_cdecl_def)
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   461
    apply (drule bspec, assumption)
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   462
    apply (clarsimp simp add: wf_mdecl_def)
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   463
    apply (drule bspec, assumption)
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   464
    apply (clarsimp simp add: make_Cert_def)
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   465
    apply (clarsimp dest!: wtl_method_complete)    
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   466
    done
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   467
10628
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   468
qed   
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   469
      
9559
kleing
parents: 9549
diff changeset
   470
lemmas [simp] = split_paired_Ex
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   471
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   472
end