src/HOL/MicroJava/Comp/CorrComp.thy
author nipkow
Tue, 14 Mar 2023 18:19:10 +0100
changeset 77645 7edbb16bc60f
parent 67613 ce654b0e6d69
child 80914 d97fdabd9e2b
permissions -rw-r--r--
Adjusted to new map update priorities
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/Comp/CorrComp.thy
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     2
    Author:     Martin Strecker
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     3
*)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     4
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     5
(* Compiler correctness statement and proof *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     6
15860
a344c4284972 partial modernising of theory headers
paulson
parents: 15481
diff changeset
     7
theory CorrComp
65538
a39ef48fbee0 tuned imports;
wenzelm
parents: 60304
diff changeset
     8
imports "../J/JTypeSafe" LemmasComp
15860
a344c4284972 partial modernising of theory headers
paulson
parents: 15481
diff changeset
     9
begin
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    10
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
    11
declare wf_prog_ws_prog [simp add]
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    12
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    13
(* If no exception is present after evaluation/execution, 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    14
  none can have been present before *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    15
lemma eval_evals_exec_xcpt:
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 20432
diff changeset
    16
 "(G \<turnstile> xs -ex\<succ>val-> xs' \<longrightarrow> gx xs' = None \<longrightarrow> gx xs = None) \<and>
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 20432
diff changeset
    17
  (G \<turnstile> xs -exs[\<succ>]vals-> xs' \<longrightarrow> gx xs' = None \<longrightarrow> gx xs = None) \<and>
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 20432
diff changeset
    18
  (G \<turnstile> xs -st-> xs' \<longrightarrow> gx xs' = None \<longrightarrow> gx xs = None)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
    19
  by (induct rule: eval_evals_exec.induct) auto
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    20
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    21
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    22
(* instance of eval_evals_exec_xcpt for eval *)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 20432
diff changeset
    23
lemma eval_xcpt: "G \<turnstile> xs -ex\<succ>val-> xs' \<Longrightarrow> gx xs' = None \<Longrightarrow> gx xs = None"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    24
 (is "?H1 \<Longrightarrow> ?H2 \<Longrightarrow> ?T")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    25
proof-
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    26
  assume h1: ?H1
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    27
  assume h2: ?H2
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    28
  from h1 h2 eval_evals_exec_xcpt show "?T" by simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    29
qed
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    30
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    31
(* instance of eval_evals_exec_xcpt for evals *)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 20432
diff changeset
    32
lemma evals_xcpt: "G \<turnstile> xs -exs[\<succ>]vals-> xs' \<Longrightarrow> gx xs' = None \<Longrightarrow> gx xs = None"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    33
 (is "?H1 \<Longrightarrow> ?H2 \<Longrightarrow> ?T")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    34
proof-
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    35
  assume h1: ?H1
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    36
  assume h2: ?H2
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    37
  from h1 h2 eval_evals_exec_xcpt show "?T" by simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    38
qed
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    39
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    40
(* instance of eval_evals_exec_xcpt for exec *)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 20432
diff changeset
    41
lemma exec_xcpt: "G \<turnstile> xs -st-> xs' \<Longrightarrow> gx xs' = None \<Longrightarrow> gx xs = None"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    42
 (is "?H1 \<Longrightarrow> ?H2 \<Longrightarrow> ?T")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    43
proof-
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    44
  assume h1: ?H1
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    45
  assume h2: ?H2
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    46
  from h1 h2 eval_evals_exec_xcpt show "?T" by simp
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    47
qed
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    48
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    49
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    50
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    51
theorem exec_all_trans: "\<lbrakk>(exec_all G s0 s1); (exec_all G s1 s2)\<rbrakk> \<Longrightarrow> (exec_all G s0 s2)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
    52
 by (auto simp: exec_all_def elim: Transitive_Closure.rtrancl_trans)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    53
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    54
theorem exec_all_refl: "exec_all G s s"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
    55
 by (simp only: exec_all_def)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    56
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    57
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    58
theorem exec_instr_in_exec_all:
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    59
  "\<lbrakk> exec_instr i G hp stk lvars C S pc frs =  (None, hp', frs');
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    60
             gis (gmb G C S) ! pc = i\<rbrakk>  \<Longrightarrow>
53024
e0968e1f6fe9 more symbolic notation;
wenzelm
parents: 52866
diff changeset
    61
       G \<turnstile> (None, hp, (stk, lvars, C, S, pc) # frs) \<midarrow>jvm\<rightarrow> (None, hp', frs')"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
    62
  apply (simp only: exec_all_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
    63
  apply (rule rtrancl_refl [THEN rtrancl_into_rtrancl])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
    64
  apply (simp add: gis_def gmb_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
    65
  apply (case_tac frs', simp+)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
    66
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    67
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    68
theorem exec_all_one_step: "
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    69
  \<lbrakk> gis (gmb G C S) = pre @ (i # post); pc0 = length pre;
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    70
  (exec_instr i G hp0 stk0 lvars0 C S pc0 frs) = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    71
  (None, hp1, (stk1,lvars1,C,S, Suc pc0)#frs) \<rbrakk>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    72
  \<Longrightarrow> 
53024
e0968e1f6fe9 more symbolic notation;
wenzelm
parents: 52866
diff changeset
    73
  G \<turnstile> (None, hp0, (stk0,lvars0,C,S, pc0)#frs) \<midarrow>jvm\<rightarrow> 
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    74
  (None, hp1, (stk1,lvars1,C,S, Suc pc0)#frs)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
    75
  apply (unfold exec_all_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
    76
  apply (rule r_into_rtrancl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
    77
  apply (simp add: gis_def gmb_def split_beta)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
    78
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    79
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    80
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    81
(***********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    82
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33639
diff changeset
    83
definition progression :: "jvm_prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> 
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    84
                 aheap \<Rightarrow> opstack \<Rightarrow> locvars \<Rightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    85
                 bytecode \<Rightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    86
                 aheap \<Rightarrow> opstack \<Rightarrow> locvars \<Rightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    87
                 bool"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33639
diff changeset
    88
  ("{_,_,_} \<turnstile> {_, _, _} >- _ \<rightarrow> {_, _, _}" [61,61,61,61,61,61,90,61,61,61]60) where
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    89
  "{G,C,S} \<turnstile> {hp0, os0, lvars0} >- instrs \<rightarrow> {hp1, os1, lvars1} ==
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
    90
  \<forall>pre post frs.
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    91
  (gis (gmb G C S) = pre @ instrs @ post) \<longrightarrow>
53024
e0968e1f6fe9 more symbolic notation;
wenzelm
parents: 52866
diff changeset
    92
   G \<turnstile> (None,hp0,(os0,lvars0,C,S,length pre)#frs) \<midarrow>jvm\<rightarrow>
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    93
       (None,hp1,(os1,lvars1,C,S,(length pre) + (length instrs))#frs)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    94
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    95
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    96
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    97
lemma progression_call: 
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
    98
  "\<lbrakk> \<forall>pc frs.
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    99
  exec_instr instr G hp0 os0 lvars0 C S pc frs =
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   100
      (None, hp', (os', lvars', C', S', 0) # (fr pc) # frs) \<and> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   101
  gis (gmb G C' S') = instrs' @ [Return] \<and> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   102
  {G, C', S'} \<turnstile> {hp', os', lvars'} >- instrs' \<rightarrow> {hp'', os'', lvars''}  \<and>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   103
  exec_instr Return G hp'' os'' lvars'' C' S' (length instrs') 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   104
                                               ((fr pc) # frs) =
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   105
      (None, hp2, (os2, lvars2, C, S, Suc pc) # frs) \<rbrakk> \<Longrightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   106
  {G, C, S} \<turnstile> {hp0, os0, lvars0} >-[instr]\<rightarrow> {hp2,os2,lvars2}"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   107
  apply (simp only: progression_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   108
  apply (intro strip)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   109
  apply (drule_tac x="length pre" in spec)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   110
  apply (drule_tac x="frs" in spec)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   111
  apply clarify
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   112
  apply (rule exec_all_trans)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   113
   apply (rule exec_instr_in_exec_all)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   114
    apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   115
   apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   116
  apply (rule exec_all_trans)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   117
   apply (simp only: append_Nil)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   118
   apply (drule_tac x="[]" in spec)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   119
   apply (simp only: list.simps list.size)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   120
   apply blast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   121
  apply (rule exec_instr_in_exec_all)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   122
   apply auto
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   123
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   124
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   125
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   126
lemma progression_transitive: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   127
  "\<lbrakk> instrs_comb = instrs0 @ instrs1; 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   128
  {G, C, S} \<turnstile> {hp0, os0, lvars0} >- instrs0 \<rightarrow> {hp1, os1, lvars1};
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   129
  {G, C, S} \<turnstile> {hp1, os1, lvars1} >- instrs1 \<rightarrow> {hp2, os2, lvars2} \<rbrakk>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   130
  \<Longrightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   131
  {G, C, S} \<turnstile> {hp0, os0, lvars0} >- instrs_comb \<rightarrow> {hp2, os2, lvars2}"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   132
  apply (simp only: progression_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   133
  apply (intro strip)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   134
  apply (rule_tac ?s1.0 = "Norm (hp1, (os1, lvars1, C, S,
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   135
                           length pre + length instrs0) # frs)"  
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   136
               in exec_all_trans)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   137
   apply (simp only: append_assoc)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   138
  apply (erule thin_rl, erule thin_rl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   139
  apply (drule_tac x="pre @ instrs0" in spec)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   140
  apply (simp add: add.assoc)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   141
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   142
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   143
lemma progression_refl: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   144
  "{G, C, S} \<turnstile> {hp0, os0, lvars0} >- [] \<rightarrow> {hp0, os0, lvars0}"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   145
  apply (simp add: progression_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   146
  apply (intro strip)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   147
  apply (rule exec_all_refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   148
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   149
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   150
lemma progression_one_step: "
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   151
  \<forall>pc frs. 
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   152
  (exec_instr i G hp0 os0 lvars0 C S pc frs) = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   153
  (None, hp1, (os1,lvars1,C,S, Suc pc)#frs)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   154
  \<Longrightarrow> {G, C, S} \<turnstile> {hp0, os0, lvars0} >- [i] \<rightarrow> {hp1, os1, lvars1}"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   155
  apply (unfold progression_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   156
  apply (intro strip)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   157
  apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   158
  apply (rule exec_all_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   159
    apply auto
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   160
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   161
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   162
(*****)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33639
diff changeset
   163
definition jump_fwd :: "jvm_prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> 
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   164
                 aheap \<Rightarrow> locvars \<Rightarrow> opstack \<Rightarrow> opstack \<Rightarrow> 
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33639
diff changeset
   165
                 instr \<Rightarrow> bytecode \<Rightarrow> bool" where
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   166
  "jump_fwd G C S hp lvars os0 os1 instr instrs ==
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   167
  \<forall>pre post frs.
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   168
  (gis (gmb G C S) = pre @ instr # instrs @ post) \<longrightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   169
   exec_all G (None,hp,(os0,lvars,C,S, length pre)#frs)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   170
    (None,hp,(os1,lvars,C,S,(length pre) + (length instrs) + 1)#frs)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   171
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   172
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   173
lemma jump_fwd_one_step:
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   174
  "\<forall>pc frs.
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   175
  exec_instr instr G hp os0 lvars C S pc frs = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   176
    (None, hp, (os1, lvars, C, S, pc + (length instrs) + 1)#frs)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   177
  \<Longrightarrow> jump_fwd G C S hp lvars os0 os1 instr instrs"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   178
  apply (unfold jump_fwd_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   179
  apply (intro strip)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   180
  apply (rule exec_instr_in_exec_all)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   181
   apply auto
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   182
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   183
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   184
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   185
lemma jump_fwd_progression_aux: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   186
  "\<lbrakk> instrs_comb = instr # instrs0 @ instrs1; 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   187
     jump_fwd G C S hp lvars os0 os1 instr instrs0;
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   188
     {G, C, S} \<turnstile> {hp, os1, lvars} >- instrs1 \<rightarrow> {hp2, os2, lvars2} \<rbrakk> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   189
  \<Longrightarrow> {G, C, S} \<turnstile> {hp, os0, lvars} >- instrs_comb \<rightarrow> {hp2, os2, lvars2}"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   190
  apply (simp only: progression_def jump_fwd_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   191
  apply (intro strip)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   192
  apply (rule_tac ?s1.0 = "Norm(hp, (os1, lvars, C, S, length pre + length instrs0 + 1) # frs)" in exec_all_trans)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   193
   apply (simp only: append_assoc)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   194
   apply (subgoal_tac "pre @ (instr # instrs0 @ instrs1) @ post = pre @ instr # instrs0 @ (instrs1 @ post)")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   195
    apply blast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   196
   apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   197
  apply (erule thin_rl, erule thin_rl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   198
  apply (drule_tac x="pre @ instr # instrs0" in spec)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   199
  apply (simp add: add.assoc)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   200
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   201
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   202
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   203
lemma jump_fwd_progression:
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   204
  "\<lbrakk> instrs_comb = instr # instrs0 @ instrs1; 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   205
  \<forall> pc frs.
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   206
  exec_instr instr G hp os0 lvars C S pc frs = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   207
    (None, hp, (os1, lvars, C, S, pc + (length instrs0) + 1)#frs);
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   208
  {G, C, S} \<turnstile> {hp, os1, lvars} >- instrs1 \<rightarrow> {hp2, os2, lvars2} \<rbrakk> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   209
  \<Longrightarrow> {G, C, S}  \<turnstile> {hp, os0, lvars} >- instrs_comb \<rightarrow> {hp2, os2, lvars2}"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   210
  apply (rule jump_fwd_progression_aux)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   211
    apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   212
   apply (rule jump_fwd_one_step, assumption+)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   213
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   214
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   215
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   216
(* note: instrs and instr reversed wrt. jump_fwd *)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33639
diff changeset
   217
definition jump_bwd :: "jvm_prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> 
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   218
                 aheap \<Rightarrow> locvars \<Rightarrow> opstack \<Rightarrow> opstack \<Rightarrow> 
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33639
diff changeset
   219
                 bytecode \<Rightarrow> instr \<Rightarrow> bool" where
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   220
  "jump_bwd G C S hp lvars os0 os1 instrs instr ==
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   221
  \<forall> pre post frs.
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   222
  (gis (gmb G C S) = pre @ instrs @ instr # post) \<longrightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   223
   exec_all G (None,hp,(os0,lvars,C,S, (length pre) + (length instrs))#frs)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   224
    (None,hp,(os1,lvars,C,S, (length pre))#frs)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   225
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   226
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   227
lemma jump_bwd_one_step:
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   228
  "\<forall> pc frs.
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   229
  exec_instr instr G hp os0 lvars C S (pc + (length instrs)) frs = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   230
    (None, hp, (os1, lvars, C, S, pc)#frs)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   231
  \<Longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   232
  jump_bwd G C S hp lvars os0 os1 instrs instr"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   233
  apply (unfold jump_bwd_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   234
  apply (intro strip)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   235
  apply (rule exec_instr_in_exec_all)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   236
   apply auto
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   237
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   238
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   239
lemma jump_bwd_progression: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   240
  "\<lbrakk> instrs_comb = instrs @ [instr]; 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   241
  {G, C, S} \<turnstile> {hp0, os0, lvars0} >- instrs \<rightarrow> {hp1, os1, lvars1};
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   242
  jump_bwd G C S hp1 lvars1 os1 os2 instrs instr;
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   243
  {G, C, S} \<turnstile> {hp1, os2, lvars1} >- instrs_comb \<rightarrow> {hp3, os3, lvars3} \<rbrakk> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   244
  \<Longrightarrow> {G, C, S}  \<turnstile> {hp0, os0, lvars0} >- instrs_comb \<rightarrow> {hp3, os3, lvars3}"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   245
  apply (simp only: progression_def jump_bwd_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   246
  apply (intro strip)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   247
  apply (rule exec_all_trans, force)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   248
  apply (rule exec_all_trans, force)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   249
  apply (rule exec_all_trans, force)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   250
  apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   251
  apply (rule exec_all_refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   252
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   253
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   254
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   255
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   256
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   257
(* class C with signature S is defined in program G *)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33639
diff changeset
   258
definition class_sig_defined :: "'c prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> bool" where
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   259
  "class_sig_defined G C S == 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   260
  is_class G C \<and> (\<exists> D rT mb. (method (G, C) S = Some (D, rT, mb)))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   261
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   262
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   263
(* The environment of a java method body 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   264
  (characterized by class and signature) *)
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33639
diff changeset
   265
definition env_of_jmb :: "java_mb prog \<Rightarrow> cname \<Rightarrow> sig \<Rightarrow> java_mb env" where
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   266
  "env_of_jmb G C S == 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   267
  (let (mn,pTs) = S;
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   268
       (D,rT,(pns,lvars,blk,res)) = the(method (G, C) S) in
77645
7edbb16bc60f Adjusted to new map update priorities
nipkow
parents: 67613
diff changeset
   269
  (G,(map_of lvars)(pns[\<mapsto>]pTs, This\<mapsto>Class C)))"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   270
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   271
lemma env_of_jmb_fst [simp]: "fst (env_of_jmb G C S) = G"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   272
by (simp add: env_of_jmb_def split_beta)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   273
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   274
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   275
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   276
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   277
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   278
lemma method_preserves [rule_format (no_asm)]:
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   279
  "\<lbrakk> wf_prog wf_mb G; is_class G C; 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   280
  \<forall> S rT mb. \<forall> cn \<in> fst ` set G. wf_mdecl wf_mb G cn (S,rT,mb)  \<longrightarrow> (P cn S (rT,mb))\<rbrakk>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   281
 \<Longrightarrow> \<forall> D. 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   282
  method (G, C) S = Some (D, rT, mb) \<longrightarrow> (P D S (rT,mb))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   283
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   284
  apply (frule wf_prog_ws_prog [THEN wf_subcls1])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   285
  apply (rule subcls1_induct, assumption, assumption)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   286
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   287
   apply (intro strip)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   288
   apply ((drule spec)+, drule_tac x="Object" in bspec)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   289
    apply (simp add: wf_prog_def ws_prog_def wf_syscls_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   290
   apply (subgoal_tac "D=Object") apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   291
    apply (drule mp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   292
     apply (frule_tac C=Object in method_wf_mdecl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   293
       apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   294
      apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   295
     apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   296
    apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   297
   apply simp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   298
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   299
  apply (simplesubst method_rec)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   300
    apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   301
   apply force
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   302
  apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   303
  apply (simp only: map_add_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   304
  apply (split option.split)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   305
  apply (rule conjI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   306
   apply force
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   307
  apply (intro strip)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   308
  apply (frule_tac ?P1.0 = "wf_mdecl wf_mb G Ca" and
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   309
                   ?P2.0 = "%(S, (Da, rT, mb)). P Da S (rT, mb)" in map_of_map_prop)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   310
    apply (force simp: wf_cdecl_def)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   311
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   312
   apply clarify
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   313
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   314
   apply (simp only: class_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   315
   apply (drule map_of_SomeD)+
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   316
   apply (frule_tac A="set G" and f=fst in imageI, simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   317
   apply blast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   318
  apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   319
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   320
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   321
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   322
lemma method_preserves_length:
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   323
  "\<lbrakk> wf_java_prog G; is_class G C; 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   324
  method (G, C) (mn,pTs) = Some (D, rT, pns, lvars, blk, res)\<rbrakk>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   325
 \<Longrightarrow> length pns = length pTs"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   326
  apply (frule_tac P="%D (mn,pTs) (rT, pns, lvars, blk, res). length pns = length pTs"
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   327
                   in method_preserves)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   328
  apply (auto simp: wf_mdecl_def wf_java_mdecl_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   329
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   330
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   331
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   332
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33639
diff changeset
   333
definition wtpd_expr :: "java_mb env \<Rightarrow> expr \<Rightarrow> bool" where
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   334
  "wtpd_expr E e == (\<exists> T. E\<turnstile>e :: T)"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33639
diff changeset
   335
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33639
diff changeset
   336
definition wtpd_exprs :: "java_mb env \<Rightarrow> (expr list) \<Rightarrow> bool" where
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   337
  "wtpd_exprs E e == (\<exists> T. E\<turnstile>e [::] T)"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33639
diff changeset
   338
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33639
diff changeset
   339
definition wtpd_stmt :: "java_mb env \<Rightarrow> stmt \<Rightarrow> bool" where
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   340
  "wtpd_stmt E c == (E\<turnstile>c \<surd>)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   341
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   342
lemma wtpd_expr_newc: "wtpd_expr E (NewC C) \<Longrightarrow> is_class (prg E) C"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   343
  by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   344
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   345
lemma wtpd_expr_cast: "wtpd_expr E (Cast cn e) \<Longrightarrow> (wtpd_expr E e)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   346
  by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   347
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   348
lemma wtpd_expr_lacc:
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   349
  "\<lbrakk> wtpd_expr (env_of_jmb G C S) (LAcc vn); class_sig_defined G C S \<rbrakk>
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   350
  \<Longrightarrow> vn \<in> set (gjmb_plns (gmb G C S)) \<or> vn = This"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   351
  apply (clarsimp simp: wtpd_expr_def env_of_jmb_def class_sig_defined_def galldefs)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   352
  apply (case_tac S)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   353
  apply (erule ty_expr.cases; fastforce dest: map_upds_SomeD map_of_SomeD fst_in_set_lemma)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   354
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   355
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   356
lemma wtpd_expr_lass: "wtpd_expr E (vn::=e)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   357
  \<Longrightarrow> (vn \<noteq> This) & (wtpd_expr E (LAcc vn)) & (wtpd_expr E e)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   358
  by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   359
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   360
lemma wtpd_expr_facc: "wtpd_expr E ({fd}a..fn) 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   361
  \<Longrightarrow> (wtpd_expr E a)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   362
  by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   363
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   364
lemma wtpd_expr_fass: "wtpd_expr E ({fd}a..fn:=v) 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   365
  \<Longrightarrow> (wtpd_expr E ({fd}a..fn)) & (wtpd_expr E v)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   366
  by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   367
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   368
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   369
lemma wtpd_expr_binop: "wtpd_expr E (BinOp bop e1 e2)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   370
  \<Longrightarrow> (wtpd_expr E e1) & (wtpd_expr E e2)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   371
  by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   372
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   373
lemma wtpd_exprs_cons: "wtpd_exprs E (e # es)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   374
  \<Longrightarrow> (wtpd_expr E e) & (wtpd_exprs E es)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   375
  by (simp only: wtpd_exprs_def wtpd_expr_def, erule exE, erule ty_exprs.cases, auto)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   376
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   377
lemma wtpd_stmt_expr: "wtpd_stmt E (Expr e) \<Longrightarrow> (wtpd_expr E e)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   378
  by (simp only: wtpd_stmt_def wtpd_expr_def, erule wt_stmt.cases, auto)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   379
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   380
lemma wtpd_stmt_comp: "wtpd_stmt E (s1;; s2) \<Longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   381
   (wtpd_stmt E s1) &  (wtpd_stmt E s2)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   382
   by (simp only: wtpd_stmt_def wtpd_expr_def, erule wt_stmt.cases, auto)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   383
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   384
lemma wtpd_stmt_cond: "wtpd_stmt E (If(e) s1 Else s2) \<Longrightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   385
   (wtpd_expr E e) & (wtpd_stmt E s1) &  (wtpd_stmt E s2)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   386
  & (E\<turnstile>e::PrimT Boolean)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   387
  by (simp only: wtpd_stmt_def wtpd_expr_def, erule wt_stmt.cases, auto)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   388
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   389
lemma wtpd_stmt_loop: "wtpd_stmt E (While(e) s) \<Longrightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   390
   (wtpd_expr E e) & (wtpd_stmt E s) & (E\<turnstile>e::PrimT Boolean)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   391
   by (simp only: wtpd_stmt_def wtpd_expr_def, erule wt_stmt.cases, auto)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   392
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   393
lemma wtpd_expr_call: "wtpd_expr E ({C}a..mn({pTs'}ps))
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   394
  \<Longrightarrow> (wtpd_expr E a) & (wtpd_exprs E ps) 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   395
  & (length ps = length pTs') & (E\<turnstile>a::Class C)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   396
  & (\<exists> pTs md rT. 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   397
       E\<turnstile>ps[::]pTs & max_spec (prg E) C (mn, pTs) = {((md,rT),pTs')})"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   398
  apply (simp only: wtpd_expr_def wtpd_exprs_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   399
  apply (erule exE)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   400
  apply (ind_cases "E \<turnstile> {C}a..mn( {pTs'}ps) :: T" for T)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   401
  apply (auto simp: max_spec_preserves_length)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   402
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   403
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   404
lemma wtpd_blk: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   405
  "\<lbrakk> method (G, D) (md, pTs) = Some (D, rT, (pns, lvars, blk, res)); 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   406
  wf_prog wf_java_mdecl G; is_class G D \<rbrakk>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   407
 \<Longrightarrow> wtpd_stmt (env_of_jmb G D (md, pTs)) blk"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   408
  apply (simp add: wtpd_stmt_def env_of_jmb_def)
77645
7edbb16bc60f Adjusted to new map update priorities
nipkow
parents: 67613
diff changeset
   409
  apply (frule_tac P="%D (md, pTs) (rT, (pns, lvars, blk, res)). (G, (map_of lvars)(pns[\<mapsto>]pTs, This\<mapsto>Class D)) \<turnstile> blk \<surd> " in method_preserves)
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   410
     apply (auto simp: wf_mdecl_def wf_java_mdecl_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   411
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   412
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   413
lemma wtpd_res: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   414
  "\<lbrakk> method (G, D) (md, pTs) = Some (D, rT, (pns, lvars, blk, res)); 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   415
  wf_prog wf_java_mdecl G; is_class G D \<rbrakk>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   416
 \<Longrightarrow> wtpd_expr (env_of_jmb G D (md, pTs)) res"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   417
  apply (simp add: wtpd_expr_def env_of_jmb_def)
77645
7edbb16bc60f Adjusted to new map update priorities
nipkow
parents: 67613
diff changeset
   418
  apply (frule_tac P="%D (md, pTs) (rT, (pns, lvars, blk, res)). \<exists>T. (G, (map_of lvars)(pns[\<mapsto>]pTs, This\<mapsto>Class D)) \<turnstile> res :: T " in method_preserves)
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   419
     apply (auto simp: wf_mdecl_def wf_java_mdecl_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   420
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   421
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   422
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   423
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   424
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   425
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   426
(* Is there a more elegant proof? *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   427
lemma evals_preserves_length:
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   428
  "G\<turnstile> xs -es[\<succ>]vs-> (None, s) \<Longrightarrow> length es = length vs"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   429
  apply (subgoal_tac 
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   430
    "\<forall> xs'. (G \<turnstile> xk -xj\<succ>xi-> xh \<longrightarrow> True) & 
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   431
    (G\<turnstile> xs -es[\<succ>]vs-> xs' \<longrightarrow>  (\<exists> s. (xs' = (None, s))) \<longrightarrow> 
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   432
    length es = length vs) &
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   433
    (G \<turnstile> xc -xb-> xa \<longrightarrow> True)")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   434
   apply blast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   435
  apply (rule allI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   436
  apply (rule Eval.eval_evals_exec.induct; simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   437
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   438
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   439
(***********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   440
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   441
(* required for translation of BinOp *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   442
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   443
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   444
lemma progression_Eq : "{G, C, S} \<turnstile>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   445
  {hp, (v2 # v1 # os), lvars} 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   446
  >- [Ifcmpeq 3, LitPush (Bool False), Goto 2, LitPush (Bool True)] \<rightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   447
  {hp, (Bool (v1 = v2) # os), lvars}"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   448
  apply (case_tac "v1 = v2")
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   449
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   450
   (* case v1 = v2 *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   451
   apply (rule_tac ?instrs1.0 = "[LitPush (Bool True)]" in jump_fwd_progression)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   452
     apply (auto simp: nat_add_distrib)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   453
   apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   454
   apply simp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   455
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   456
  (* case v1 \<noteq> v2 *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   457
  apply (rule progression_one_step [THEN HOL.refl [THEN progression_transitive], simplified])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   458
   apply auto
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   459
  apply (rule progression_one_step [THEN HOL.refl [THEN progression_transitive], simplified])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   460
   apply auto
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   461
  apply (rule_tac ?instrs1.0 = "[]" in jump_fwd_progression)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   462
    apply (auto simp: nat_add_distrib intro: progression_refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   463
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   464
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   465
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   466
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   467
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   468
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   469
(* to avoid automatic pair splits *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   470
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   471
declare split_paired_All [simp del] split_paired_Ex [simp del]
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   472
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   473
lemma distinct_method:
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   474
  "\<lbrakk> wf_java_prog G; is_class G C; method (G, C) S = Some (D, rT, pns, lvars, blk, res) \<rbrakk> \<Longrightarrow> 
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   475
  distinct (gjmb_plns (gmb G C S))"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   476
  apply (frule method_wf_mdecl [THEN conjunct2],  assumption, assumption)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   477
  apply (case_tac S)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   478
  apply (simp add: wf_mdecl_def wf_java_mdecl_def galldefs)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   479
  apply (simp add: unique_def map_of_in_set)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   480
  apply blast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   481
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   482
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   483
lemma distinct_method_if_class_sig_defined : 
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   484
  "\<lbrakk> wf_java_prog G; class_sig_defined G C S \<rbrakk> \<Longrightarrow> distinct (gjmb_plns (gmb G C S))"
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   485
  by (auto intro: distinct_method simp: class_sig_defined_def)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   486
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   487
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   488
lemma method_yields_wf_java_mdecl: "\<lbrakk> wf_java_prog G; is_class G C;
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   489
  method (G, C) S = Some (D, rT, pns, lvars, blk, res) \<rbrakk>  \<Longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   490
  wf_java_mdecl G D (S,rT,(pns,lvars,blk,res))"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   491
  apply (frule method_wf_mdecl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   492
    apply (auto simp: wf_mdecl_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   493
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   494
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   495
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   496
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   497
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   498
lemma progression_lvar_init_aux [rule_format (no_asm)]: "
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   499
  \<forall> zs prfx lvals lvars0. 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   500
  lvars0 =  (zs @ lvars) \<longrightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   501
  (disjoint_varnames pns lvars0 \<longrightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   502
  (length lvars = length lvals) \<longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   503
  (Suc(length pns + length zs) = length prfx) \<longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   504
   ({cG, D, S} \<turnstile> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   505
    {h, os, (prfx @ lvals)}
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   506
    >- (concat (map (compInit (pns, lvars0, blk, res)) lvars)) \<rightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   507
    {h, os, (prfx @ (map (\<lambda>p. (default_val (snd p))) lvars))}))"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   508
  apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   509
  apply (induct lvars)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   510
   apply (clarsimp, rule progression_refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   511
  apply (intro strip)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   512
  apply (case_tac lvals)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   513
   apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   514
  apply (simp (no_asm_simp) )
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   515
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   516
  apply (rule_tac ?lvars1.0 = "(prfx @ [default_val (snd a)]) @ list" in progression_transitive, rule HOL.refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   517
   apply (case_tac a)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   518
   apply (simp (no_asm_simp) add: compInit_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   519
   apply (rule_tac ?instrs0.0 = "[load_default_val b]" in progression_transitive, simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   520
    apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   521
    apply (simp (no_asm_simp) add: load_default_val_def)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   522
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   523
   apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   524
   apply (simp (no_asm_simp))
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   525
   apply (rule conjI, simp)+
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   526
   apply (simp add: index_of_var2)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   527
  apply (drule_tac x="zs @ [a]" in spec) (* instantiate zs *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   528
  apply (drule mp, simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   529
  apply (drule_tac x="(prfx @ [default_val (snd a)])" in spec) (* instantiate prfx *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   530
  apply auto
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   531
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   532
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   533
lemma progression_lvar_init [rule_format (no_asm)]: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   534
  "\<lbrakk> wf_java_prog G; is_class G C;
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   535
  method (G, C) S = Some (D, rT, pns, lvars, blk, res) \<rbrakk> \<Longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   536
  length pns = length pvs \<longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   537
  (\<forall> lvals. 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   538
  length lvars = length lvals \<longrightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   539
   {cG, D, S} \<turnstile>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   540
   {h, os, (a' # pvs @ lvals)}
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   541
   >- (compInitLvars (pns, lvars, blk, res) lvars) \<rightarrow>
77645
7edbb16bc60f Adjusted to new map update priorities
nipkow
parents: 67613
diff changeset
   542
   {h, os, (locvars_xstate G C S (Norm (h, (init_vars lvars)(pns[\<mapsto>]pvs, This\<mapsto>a'))))})"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   543
  apply (simp only: compInitLvars_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   544
  apply (frule method_yields_wf_java_mdecl, assumption, assumption)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   545
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   546
  apply (simp only: wf_java_mdecl_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   547
  apply (subgoal_tac "(\<forall>y\<in>set pns. y \<notin> set (map fst lvars))")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   548
   apply (simp add: init_vars_def locvars_xstate_def locvars_locals_def galldefs unique_def split_def map_of_map_as_map_upd del: map_map)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   549
   apply (intro strip)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   550
   apply (simp (no_asm_simp) only: append_Cons [symmetric])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   551
   apply (rule progression_lvar_init_aux)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   552
      apply (auto simp: unique_def map_of_in_set disjoint_varnames_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   553
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   554
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   555
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   556
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   557
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   558
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   559
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   560
lemma state_ok_eval:
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   561
  "\<lbrakk>xs::\<preceq>E; wf_java_prog (prg E); wtpd_expr E e; (prg E)\<turnstile>xs -e\<succ>v -> xs'\<rbrakk> \<Longrightarrow>  xs'::\<preceq>E"
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   562
  apply (simp only: wtpd_expr_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   563
  apply (erule exE)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   564
  apply (case_tac xs', case_tac xs)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   565
  apply (auto intro: eval_type_sound [THEN conjunct1])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   566
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   567
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   568
lemma state_ok_evals:
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   569
  "\<lbrakk>xs::\<preceq>E; wf_java_prog (prg E); wtpd_exprs E es; prg E \<turnstile> xs -es[\<succ>]vs-> xs'\<rbrakk> \<Longrightarrow> xs'::\<preceq>E"
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   570
  apply (simp only: wtpd_exprs_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   571
  apply (erule exE)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   572
  apply (case_tac xs)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   573
  apply (case_tac xs')
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   574
  apply (auto intro: evals_type_sound [THEN conjunct1])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   575
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   576
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   577
lemma state_ok_exec:
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   578
  "\<lbrakk>xs::\<preceq>E; wf_java_prog (prg E); wtpd_stmt E st; prg E \<turnstile> xs -st-> xs'\<rbrakk> \<Longrightarrow>  xs'::\<preceq>E"
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   579
  apply (simp only: wtpd_stmt_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   580
  apply (case_tac xs', case_tac xs)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   581
  apply (auto dest: exec_type_sound)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   582
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   583
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   584
lemma state_ok_init: 
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   585
  "\<lbrakk> wf_java_prog G; (x, h, l)::\<preceq>(env_of_jmb G C S); 
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   586
  is_class G dynT;
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   587
  method (G, dynT) (mn, pTs) = Some (md, rT, pns, lvars, blk, res);
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   588
  list_all2 (conf G h) pvs pTs; G,h \<turnstile> a' ::\<preceq> Class md\<rbrakk>
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   589
  \<Longrightarrow>
77645
7edbb16bc60f Adjusted to new map update priorities
nipkow
parents: 67613
diff changeset
   590
  (np a' x, h, (init_vars lvars)(pns[\<mapsto>]pvs, This\<mapsto>a'))::\<preceq>(env_of_jmb G md (mn, pTs))"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   591
  apply (frule wf_prog_ws_prog)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   592
  apply (frule method_in_md [THEN conjunct2], assumption+)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   593
  apply (frule method_yields_wf_java_mdecl, assumption, assumption)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   594
  apply (simp add: env_of_jmb_def gs_def conforms_def split_beta)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   595
  apply (simp add: wf_java_mdecl_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   596
  apply (rule conjI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   597
   apply (rule lconf_ext)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   598
    apply (rule lconf_ext_list)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   599
       apply (rule lconf_init_vars)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   600
       apply (auto dest: Ball_set_table)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   601
  apply (simp add: np_def xconf_raise_if)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   602
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   603
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   604
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   605
lemma ty_exprs_list_all2 [rule_format (no_asm)]: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   606
  "(\<forall> Ts. (E \<turnstile> es [::] Ts) = list_all2 (\<lambda>e T. E \<turnstile> e :: T) es Ts)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   607
  apply (rule list.induct)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   608
   apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   609
   apply (rule allI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   610
   apply (rule iffI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   611
    apply (ind_cases "E \<turnstile> [] [::] Ts" for Ts, assumption)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   612
   apply simp apply (rule WellType.Nil)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   613
  apply (simp add: list_all2_Cons1)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   614
  apply (rule allI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   615
  apply (rule iffI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   616
   apply (rename_tac a exs Ts)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   617
   apply (ind_cases "E \<turnstile> a # exs  [::] Ts" for a exs Ts) apply blast
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   618
  apply (auto intro: WellType.Cons)
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   619
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   620
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   621
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   622
lemma conf_bool: "G,h \<turnstile> v::\<preceq>PrimT Boolean \<Longrightarrow> \<exists> b. v = Bool b"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   623
  apply (simp add: conf_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   624
  apply (erule exE)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   625
  apply (case_tac v)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   626
      apply (auto elim: widen.cases)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   627
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   628
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   629
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   630
lemma max_spec_widen: "max_spec G C (mn, pTs) = {((md,rT),pTs')} \<Longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   631
  list_all2 (\<lambda> T T'. G \<turnstile> T \<preceq> T') pTs pTs'"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   632
  by (blast dest: singleton_in_set max_spec2appl_meths appl_methsD)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   633
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   634
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   635
lemma eval_conf: "\<lbrakk>G \<turnstile> s -e\<succ>v-> s'; wf_java_prog G; s::\<preceq>E;
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   636
  E\<turnstile>e::T; gx s' = None; prg E = G \<rbrakk> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   637
  \<Longrightarrow> G,gh s'\<turnstile>v::\<preceq>T"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   638
  apply (simp add: gh_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   639
  apply (rule_tac x3="fst s" and s3="snd s"and x'3="fst s'"
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   640
               in eval_type_sound [THEN conjunct2 [THEN conjunct1 [THEN mp]], simplified])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   641
       apply assumption+
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   642
      apply (simp (no_asm_use))
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   643
     apply (simp only: surjective_pairing [symmetric])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   644
    apply (auto simp add: gs_def gx_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   645
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   646
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   647
lemma evals_preserves_conf:
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   648
  "\<lbrakk> G\<turnstile> s -es[\<succ>]vs-> s'; G,gh s \<turnstile> t ::\<preceq> T; E \<turnstile>es[::]Ts;
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   649
  wf_java_prog G; s::\<preceq>E; 
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   650
  prg E = G \<rbrakk> \<Longrightarrow> G,gh s' \<turnstile> t ::\<preceq> T"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   651
  apply (subgoal_tac "gh s\<le>| gh s'")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   652
   apply (frule conf_hext, assumption, assumption)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   653
  apply (frule eval_evals_exec_type_sound [THEN conjunct2 [THEN conjunct1 [THEN mp]]])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   654
   apply (subgoal_tac "G \<turnstile> (gx s, (gh s, gl s)) -es[\<succ>]vs-> (gx s', (gh s', gl s'))")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   655
    apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   656
   apply (simp add: gx_def gh_def gl_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   657
  apply (case_tac E)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   658
  apply (simp add: gx_def gs_def gh_def gl_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   659
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   660
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   661
lemma eval_of_class:
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   662
  "\<lbrakk> G \<turnstile> s -e\<succ>a'-> s'; E \<turnstile> e :: Class C; wf_java_prog G; s::\<preceq>E; gx s'=None; a' \<noteq> Null; G=prg E\<rbrakk>
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   663
  \<Longrightarrow> (\<exists> lc. a' = Addr lc)"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   664
  apply (case_tac s, case_tac s', simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   665
  apply (frule eval_type_sound, (simp add: gs_def)+)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   666
  apply (case_tac a')
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   667
      apply (auto simp: conf_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   668
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   669
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   670
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   671
lemma dynT_subcls: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   672
  "\<lbrakk> a' \<noteq> Null; G,h\<turnstile>a'::\<preceq> Class C; dynT = fst (the (h (the_Addr a')));
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   673
  is_class G dynT; ws_prog G \<rbrakk> \<Longrightarrow> G\<turnstile>dynT \<preceq>C C"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   674
  apply (case_tac "C = Object")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   675
   apply (simp, rule subcls_C_Object, assumption+)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   676
  apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   677
  apply (frule non_np_objD, auto)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   678
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   679
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   680
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   681
lemma method_defined: "\<lbrakk> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   682
  m = the (method (G, dynT) (mn, pTs)); 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   683
  dynT = fst (the (h a)); is_class G dynT; wf_java_prog G; 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   684
  a' \<noteq> Null; G,h\<turnstile>a'::\<preceq> Class C; a = the_Addr a';
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   685
  \<exists>pTsa md rT. max_spec G C (mn, pTsa) = {((md, rT), pTs)} \<rbrakk>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   686
\<Longrightarrow> (method (G, dynT) (mn, pTs)) = Some m"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   687
  apply (erule exE)+
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   688
  apply (drule singleton_in_set, drule max_spec2appl_meths)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   689
  apply (simp add: appl_methds_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   690
  apply (elim exE conjE)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   691
  apply (drule widen_methd)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   692
  apply (auto intro!: dynT_subcls)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   693
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   694
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   695
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   696
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   697
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   698
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   699
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   700
(* 1. any difference between locvars_xstate \<dots> and L ??? *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   701
(* 2. possibly skip env_of_jmb ??? *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   702
theorem compiler_correctness: 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   703
  "wf_java_prog G \<Longrightarrow>
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 20432
diff changeset
   704
  (G \<turnstile> xs -ex\<succ>val-> xs' \<longrightarrow>
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   705
  gx xs = None \<longrightarrow> gx xs' = None \<longrightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   706
  (\<forall> os CL S.
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   707
  (class_sig_defined G CL S) \<longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   708
  (wtpd_expr (env_of_jmb G CL S) ex) \<longrightarrow>
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   709
  (xs ::\<preceq> (env_of_jmb G CL S)) \<longrightarrow>
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   710
  ( {TranslComp.comp G, CL, S} \<turnstile>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   711
    {gh xs, os, (locvars_xstate G CL S xs)}
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   712
    >- (compExpr (gmb G CL S) ex) \<rightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   713
    {gh xs', val#os, locvars_xstate G CL S xs'}))) \<and> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   714
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 20432
diff changeset
   715
 (G \<turnstile> xs -exs[\<succ>]vals-> xs' \<longrightarrow>
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   716
  gx xs = None \<longrightarrow> gx xs' = None \<longrightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   717
  (\<forall> os CL S.
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   718
  (class_sig_defined G CL S) \<longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   719
  (wtpd_exprs (env_of_jmb G CL S) exs) \<longrightarrow>
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   720
  (xs::\<preceq>(env_of_jmb G CL S)) \<longrightarrow>
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   721
  ( {TranslComp.comp G, CL, S} \<turnstile>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   722
    {gh xs, os, (locvars_xstate G CL S xs)}
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   723
    >- (compExprs (gmb G CL S) exs) \<rightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   724
    {gh xs', (rev vals)@os, (locvars_xstate G CL S xs')}))) \<and> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   725
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 20432
diff changeset
   726
  (G \<turnstile> xs -st-> xs' \<longrightarrow>
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   727
   gx xs = None \<longrightarrow> gx xs' = None \<longrightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   728
  (\<forall> os CL S.
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   729
  (class_sig_defined G CL S) \<longrightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   730
  (wtpd_stmt (env_of_jmb G CL S) st) \<longrightarrow>
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
   731
  (xs::\<preceq>(env_of_jmb G CL S)) \<longrightarrow>
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   732
  ( {TranslComp.comp G, CL, S} \<turnstile>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   733
    {gh xs, os, (locvars_xstate G CL S xs)}
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   734
    >- (compStmt (gmb G CL S) st) \<rightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   735
    {gh xs', os, (locvars_xstate G CL S xs')})))"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   736
  apply (rule Eval.eval_evals_exec.induct)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   737
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   738
                     (* case XcptE *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   739
                     apply simp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   740
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   741
                    (* case NewC *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   742
                    apply clarify
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 65538
diff changeset
   743
                    apply (frule wf_prog_ws_prog [THEN wf_subcls1]) (* establish  wf ((subcls1 G)\<inverse>) *)
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   744
                    apply (simp add: c_hupd_hp_invariant)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   745
                    apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   746
                    apply (rotate_tac 1, drule sym) (* reverse equation (a, None) = new_Addr (fst s) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   747
                    apply (simp add: locvars_xstate_def locvars_locals_def comp_fields)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   748
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   749
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   750
                   (* case Cast *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   751
                   apply (intro allI impI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   752
                   apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   753
                   apply (frule raise_if_NoneD)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   754
                   apply (frule wtpd_expr_cast)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   755
                   apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   756
                   apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   757
                    apply blast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   758
                   apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   759
                   apply (simp add: raise_system_xcpt_def  gh_def comp_cast_ok)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   760
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   761
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   762
                  (* case Lit *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   763
                  apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   764
                  apply (intro strip)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   765
                  apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   766
                  apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   767
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   768
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   769
                 (* case BinOp *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   770
                 apply (intro allI impI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   771
                 apply (frule_tac xs=s1 in eval_xcpt, assumption) (* establish (gx s1 = None) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   772
                 apply (frule wtpd_expr_binop)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   773
                 (* establish (s1::\<preceq> \<dots>) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   774
                 apply (frule_tac e=e1 in state_ok_eval) apply (simp (no_asm_simp))
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   775
                   apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   776
                  apply (simp (no_asm_use) only: env_of_jmb_fst)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   777
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   778
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   779
                 apply (simp (no_asm_use) only: compExpr.simps compExprs.simps)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   780
                 apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e1" in progression_transitive, simp) apply blast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   781
                 apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e2" in progression_transitive, simp) apply blast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   782
                 apply (case_tac bop)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   783
                  (*subcase bop = Eq *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   784
                  apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   785
                  apply (rule progression_Eq)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   786
                 (*subcase bop = Add *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   787
                 apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   788
                 apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   789
                 apply simp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   790
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   791
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   792
                (* case LAcc *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   793
                apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   794
                apply (intro strip)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   795
                apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   796
                apply (simp add: locvars_xstate_def locvars_locals_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   797
                apply (frule wtpd_expr_lacc)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   798
                 apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   799
                apply (simp add: gl_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   800
                apply (erule select_at_index)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   801
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   802
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   803
               (* case LAss *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   804
               apply (intro allI impI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   805
               apply (frule wtpd_expr_lass, erule conjE, erule conjE)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   806
               apply simp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   807
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   808
               apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   809
                apply blast
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   810
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   811
               apply (rule_tac ?instrs0.0 = "[Dup]" in progression_transitive, simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   812
                apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   813
                apply (simp add: gh_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   814
               apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   815
               apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   816
               apply (simp add: gh_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   817
               (* the following falls out of the general scheme *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   818
               apply (frule wtpd_expr_lacc) apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   819
               apply (rule update_at_index)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   820
                 apply (rule distinct_method_if_class_sig_defined)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   821
                  apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   822
                 apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   823
                apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   824
               apply assumption
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   825
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   826
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   827
              (* case FAcc *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   828
              apply (intro allI impI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   829
              (* establish x1 = None \<and> a' \<noteq> Null *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   830
              apply (simp (no_asm_use) only: gx_conv, frule np_NoneD)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   831
              apply (frule wtpd_expr_facc)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   832
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   833
              apply (simp (no_asm_use) only: compExpr.simps compExprs.simps)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   834
              apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   835
               apply blast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   836
              apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   837
              apply (simp add: gh_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   838
              apply (case_tac "(the (fst s1 (the_Addr a')))")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   839
              apply (simp add: raise_system_xcpt_def)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   840
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   841
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   842
             (* case FAss *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   843
             apply (intro allI impI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   844
             apply (frule wtpd_expr_fass) apply (erule conjE) apply (frule wtpd_expr_facc)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   845
             apply (simp only: c_hupd_xcpt_invariant) (* establish x2 = None *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   846
             (* establish x1 = None  and  a' \<noteq> Null  *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   847
             apply (frule_tac xs="(np a' x1, s1)" in eval_xcpt)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   848
              apply (simp only: gx_conv, simp only: gx_conv, frule np_NoneD, erule conjE)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   849
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   850
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   851
             (* establish ((Norm s1)::\<preceq> \<dots>) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   852
             apply (frule_tac e=e1 in state_ok_eval)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   853
                apply (simp (no_asm_simp) only: env_of_jmb_fst)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   854
               apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   855
              apply (simp (no_asm_use) only: env_of_jmb_fst)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   856
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   857
             apply (simp only: compExpr.simps compExprs.simps)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   858
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   859
             apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e1)" in progression_transitive, rule HOL.refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   860
              apply fast (* blast does not seem to work - why? *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   861
             apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e2)" in progression_transitive, rule HOL.refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   862
              apply fast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   863
             apply (rule_tac ?instrs0.0 = "[Dup_x1]" and ?instrs1.0 = "[Putfield fn T]" in progression_transitive, simp)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   864
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   865
              (* Dup_x1 *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   866
              apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   867
              apply (simp add: gh_def)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   868
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   869
             (* Putfield \<longrightarrow> still looks nasty*)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   870
             apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   871
             apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   872
             apply (case_tac "(the (fst s2 (the_Addr a')))")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   873
             apply (simp add: c_hupd_hp_invariant)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   874
             apply (case_tac s2)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   875
             apply (simp add: c_hupd_conv raise_system_xcpt_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   876
             apply (rule locvars_xstate_par_dep, rule HOL.refl)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   877
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   878
            defer (* method call *)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   879
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   880
            (* case XcptEs *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   881
            apply simp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   882
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   883
           (* case Nil *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   884
           apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   885
           apply (intro strip)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   886
           apply (rule progression_refl)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   887
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   888
          (* case Cons *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   889
          apply (intro allI impI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   890
          apply (frule_tac xs=s1 in evals_xcpt, simp only: gx_conv) (* establish gx s1 = None *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   891
          apply (frule wtpd_exprs_cons)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   892
          (* establish ((Norm s0)::\<preceq> \<dots>) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   893
          apply (frule_tac e=e in state_ok_eval)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   894
             apply (simp (no_asm_simp) only: env_of_jmb_fst)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   895
            apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   896
           apply (simp (no_asm_use) only: env_of_jmb_fst)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   897
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   898
          apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   899
          apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   900
           apply fast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   901
          apply fast
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   902
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   903
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   904
         (* case Statement: exception *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   905
         apply simp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   906
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   907
        (* case Skip *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   908
        apply (intro allI impI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   909
        apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   910
        apply (rule progression_refl)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   911
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   912
       (* case Expr *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   913
       apply (intro allI impI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   914
       apply (frule wtpd_stmt_expr)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   915
       apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   916
       apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   917
        apply fast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   918
       apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   919
       apply simp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   920
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   921
      (* case Comp *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   922
      apply (intro allI impI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   923
      apply (frule_tac xs=s1 in exec_xcpt, assumption) (* establish (gx s1 = None) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   924
      apply (frule wtpd_stmt_comp)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   925
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   926
      (* establish (s1::\<preceq> \<dots>) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   927
      apply (frule_tac st=c1 in state_ok_exec)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   928
         apply (simp (no_asm_simp) only: env_of_jmb_fst)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   929
        apply simp apply (simp (no_asm_use) only: env_of_jmb_fst)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   930
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   931
      apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   932
      apply (rule_tac ?instrs0.0 = "(compStmt (gmb G CL S) c1)" in progression_transitive, rule HOL.refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   933
       apply fast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   934
      apply fast
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   935
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   936
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   937
     (* case Cond *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   938
     apply (intro allI impI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   939
     apply (frule_tac xs=s1 in exec_xcpt, assumption) (* establish (gx s1 = None) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   940
     apply (frule wtpd_stmt_cond, (erule conjE)+)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   941
     (* establish (s1::\<preceq> \<dots>) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   942
     apply (frule_tac e=e in state_ok_eval)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   943
        apply (simp (no_asm_simp) only: env_of_jmb_fst)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   944
       apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   945
      apply (simp (no_asm_use) only: env_of_jmb_fst)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   946
     (* establish G,gh s1\<turnstile>v::\<preceq>PrimT Boolean *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   947
     apply (frule eval_conf, assumption+, rule env_of_jmb_fst)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   948
     apply (frule conf_bool) (* establish \<exists>b. v = Bool b *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   949
     apply (erule exE)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   950
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   951
     apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   952
     apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp))
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   953
      apply (rule progression_one_step,  simp)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   954
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   955
     apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   956
      apply fast
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   957
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   958
     apply (case_tac b)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   959
      (* case b= True *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   960
      apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   961
      apply (rule_tac ?instrs0.0 = "[Ifcmpeq (2 + int (length (compStmt (gmb G CL S) c1)))]" in progression_transitive, simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   962
       apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   963
       apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   964
      apply (rule_tac ?instrs0.0 = "(compStmt (gmb G CL S) c1)" in progression_transitive, simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   965
       apply fast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   966
      apply (rule_tac ?instrs1.0 = "[]" in jump_fwd_progression)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   967
        apply (simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   968
       apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   969
       apply (rule conjI, simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   970
       apply (simp add: nat_add_distrib)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   971
      apply (rule progression_refl)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   972
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   973
     (* case b= False *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   974
     apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   975
     apply (rule_tac ?instrs1.0 = "compStmt (gmb G CL S) c2" in jump_fwd_progression)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   976
       apply (simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   977
      apply (simp, rule conjI, rule HOL.refl, simp add: nat_add_distrib)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   978
     apply fast
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   979
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   980
    (* case exit Loop *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   981
    apply (intro allI impI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   982
    apply (frule wtpd_stmt_loop, (erule conjE)+)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   983
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   984
    (* establish G,gh s1\<turnstile>v::\<preceq>PrimT Boolean *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   985
    apply (frule eval_conf, assumption+, rule env_of_jmb_fst)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   986
    apply (frule conf_bool) (* establish \<exists>b. v = Bool b *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   987
    apply (erule exE)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   988
    apply (case_tac b)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   989
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   990
     (* case b= True \<longrightarrow> contradiction *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   991
     apply simp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   992
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   993
    (* case b= False *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   994
    apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   995
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   996
    apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp))
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   997
     apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
   998
     apply simp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   999
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1000
    apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1001
     apply fast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1002
    apply (rule_tac ?instrs1.0 = "[]" in jump_fwd_progression)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1003
      apply (simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1004
     apply (simp, rule conjI, rule HOL.refl, simp add: nat_add_distrib)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1005
    apply (rule progression_refl)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1006
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1007
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1008
   (* case continue Loop *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1009
   apply (intro allI impI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1010
   apply (frule_tac xs=s2 in exec_xcpt, assumption) (* establish (gx s2 = None) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1011
   apply (frule_tac xs=s1 in exec_xcpt, assumption) (* establish (gx s1 = None) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1012
   apply (frule wtpd_stmt_loop, (erule conjE)+)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1013
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1014
   (* establish (s1::\<preceq> \<dots>) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1015
   apply (frule_tac e=e in state_ok_eval)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1016
      apply (simp (no_asm_simp) only: env_of_jmb_fst)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1017
     apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1018
    apply (simp (no_asm_use) only: env_of_jmb_fst)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1019
   (* establish (s2::\<preceq> \<dots>) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1020
   apply (frule_tac xs=s1 and st=c in state_ok_exec)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1021
      apply (simp (no_asm_simp) only: env_of_jmb_fst)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1022
     apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1023
    apply (simp (no_asm_use) only: env_of_jmb_fst)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1024
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1025
   (* establish G,gh s1\<turnstile>v::\<preceq>PrimT Boolean *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1026
   apply (frule eval_conf, assumption+, rule env_of_jmb_fst)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1027
   apply (frule conf_bool) (* establish \<exists>b. v = Bool b *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1028
   apply (erule exE)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1029
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1030
   apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1031
   apply (rule jump_bwd_progression)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1032
      apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1033
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1034
     apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp))
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1035
      apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1036
      apply simp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1037
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1038
     apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1039
      apply fast
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1040
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1041
     apply (case_tac b)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1042
      (* case b= True *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1043
      apply simp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1044
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1045
      apply (rule_tac ?instrs0.0 = "[Ifcmpeq (2 + int (length (compStmt (gmb G CL S) c)))]" in progression_transitive, simp)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1046
       apply (rule progression_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1047
       apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1048
      apply fast
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1049
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1050
     (* case b= False \<longrightarrow> contradiction*)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1051
     apply simp
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1052
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1053
    apply (rule jump_bwd_one_step)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1054
    apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1055
   apply blast
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1056
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1057
  (*****)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1058
  (* case method call *)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1059
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1060
  apply (intro allI impI)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1061
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1062
  apply (frule_tac xs=s3 in eval_xcpt, simp only: gx_conv) (* establish gx s3 = None *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1063
  apply (frule exec_xcpt, assumption, simp (no_asm_use) only: gx_conv, frule np_NoneD) (* establish x = None \<and> a' \<noteq> Null *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1064
  apply (frule evals_xcpt, simp only: gx_conv) (* establish gx s1 = None *)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1065
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1066
  apply (frule wtpd_expr_call, (erule conjE)+)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1067
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1068
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1069
  (* assumptions about state_ok and is_class *)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1070
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1071
  (* establish s1::\<preceq> (env_of_jmb G CL S) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1072
  apply (frule_tac xs="Norm s0" and e=e in state_ok_eval)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1073
     apply (simp (no_asm_simp) only: env_of_jmb_fst, assumption, simp (no_asm_use) only: env_of_jmb_fst)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1074
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1075
  (* establish (x, h, l)::\<preceq>(env_of_jmb G CL S) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1076
  apply (frule_tac xs=s1 and xs'="(x, h, l)" in state_ok_evals)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1077
     apply (simp (no_asm_simp) only: env_of_jmb_fst, assumption, simp only: env_of_jmb_fst)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1078
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1079
  (* establish \<exists> lc. a' = Addr lc *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1080
  apply (frule (5) eval_of_class, rule env_of_jmb_fst [symmetric])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1081
  apply (subgoal_tac "G,h \<turnstile> a' ::\<preceq> Class C")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1082
   apply (subgoal_tac "is_class G dynT")
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1083
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1084
    (* establish method (G, dynT) (mn, pTs) = Some(md, rT, pns, lvars, blk, res) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1085
    apply (drule method_defined, assumption+)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1086
     apply (simp only: env_of_jmb_fst)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1087
     apply ((erule exE)+, erule conjE, (rule exI)+, assumption)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1088
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1089
    apply (subgoal_tac "is_class G md")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1090
     apply (subgoal_tac "G\<turnstile>Class dynT \<preceq> Class md")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1091
      apply (subgoal_tac " method (G, md) (mn, pTs) = Some (md, rT, pns, lvars, blk, res)")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1092
       apply (subgoal_tac "list_all2 (conf G h) pvs pTs")
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1093
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1094
        (* establish (np a' x, h, init_vars lvars(pns[\<mapsto>]pvs)(This\<mapsto>a'))::\<preceq>(env_of_jmb G md (mn, pTs)) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1095
        apply (subgoal_tac "G,h \<turnstile> a' ::\<preceq> Class dynT")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1096
         apply (frule (2) conf_widen)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1097
         apply (frule state_ok_init, assumption+)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1098
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1099
         apply (subgoal_tac "class_sig_defined G md (mn, pTs)")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1100
          apply (frule wtpd_blk, assumption, assumption)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1101
          apply (frule wtpd_res, assumption, assumption)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1102
          apply (subgoal_tac "s3::\<preceq>(env_of_jmb G md (mn, pTs))")
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1103
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1104
           apply (subgoal_tac "method (TranslComp.comp G, md) (mn, pTs) =
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1105
                               Some (md, rT, snd (snd (compMethod G md ((mn, pTs), rT, pns, lvars, blk, res))))")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1106
            prefer 2
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1107
            apply (simp add: wf_prog_ws_prog [THEN comp_method])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1108
           apply (subgoal_tac "method (TranslComp.comp G, dynT) (mn, pTs) =
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1109
                               Some (md, rT, snd (snd (compMethod G md ((mn, pTs), rT, pns, lvars, blk, res))))")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1110
            prefer 2
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1111
            apply (simp add: wf_prog_ws_prog [THEN comp_method])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1112
           apply (simp only: fst_conv snd_conv)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1113
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1114
           (* establish length pns = length pTs *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1115
           apply (frule method_preserves_length, assumption, assumption)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1116
           (* establish length pvs = length ps *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1117
           apply (frule evals_preserves_length [symmetric])
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1118
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1119
           (** start evaluating subexpressions **)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1120
           apply (simp (no_asm_use) only: compExpr.simps compExprs.simps)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1121
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1122
           (* evaluate e *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1123
           apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1124
            apply fast
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1125
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1126
           (* evaluate parameters *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1127
           apply (rule_tac ?instrs0.0 = "compExprs (gmb G CL S) ps" in progression_transitive, rule HOL.refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1128
            apply fast
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1129
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1130
           (* invokation *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1131
           apply (rule progression_call)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1132
           apply (intro allI impI conjI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1133
              (* execute Invoke statement *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1134
              apply (simp (no_asm_use) only: exec_instr.simps)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1135
              apply (erule thin_rl, erule thin_rl, erule thin_rl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1136
              apply (simp add: compMethod_def raise_system_xcpt_def)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1137
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1138
             (* get instructions of invoked method *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1139
             apply (simp (no_asm_simp) add: gis_def gmb_def compMethod_def)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1140
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1141
            (* var. initialization *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1142
            apply (rule_tac ?instrs0.0 = "(compInitLvars (pns, lvars, blk, res) lvars)" in progression_transitive, rule HOL.refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1143
             apply (rule_tac C=md in progression_lvar_init, assumption, assumption, assumption)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1144
              apply (simp (no_asm_simp)) (* length pns = length pvs *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1145
             apply (simp (no_asm_simp)) (* length lvars = length (replicate (length lvars) undefined) *)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1146
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1147
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1148
            (* body statement *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1149
            apply (rule_tac ?instrs0.0 = "compStmt (pns, lvars, blk, res) blk" in progression_transitive, rule HOL.refl)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1150
             apply (subgoal_tac "(pns, lvars, blk, res) = gmb G md (mn, pTs)")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1151
              apply (simp (no_asm_simp))
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1152
              apply (simp only: gh_conv)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1153
              apply (drule mp [OF _ TrueI])+
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1154
              apply (erule allE, erule allE, erule allE, erule impE, assumption)+
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1155
              apply ((erule impE, assumption)+, assumption)
15860
a344c4284972 partial modernising of theory headers
paulson
parents: 15481
diff changeset
  1156
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1157
             apply (simp (no_asm_use))
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1158
             apply (simp (no_asm_simp) add: gmb_def)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1159
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1160
            (* return expression *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1161
            apply (subgoal_tac "(pns, lvars, blk, res) = gmb G md (mn, pTs)")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1162
             apply (simp (no_asm_simp))
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1163
             apply (simp only: gh_conv)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1164
             apply ((drule mp, rule TrueI)+, (drule spec)+, (drule mp, assumption)+, assumption)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1165
            apply (simp (no_asm_use))
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1166
            apply (simp (no_asm_simp) add: gmb_def)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1167
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1168
           (* execute return statement *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1169
           apply (simp (no_asm_use) add: gh_def locvars_xstate_def gl_def del: drop_append)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1170
           apply (subgoal_tac "rev pvs @ a' # os = (rev (a' # pvs)) @ os")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1171
            apply (simp only: drop_append)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1172
            apply (simp (no_asm_simp))
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1173
           apply (simp (no_asm_simp))
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1174
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1175
          (* show s3::\<preceq>\<dots> *)
77645
7edbb16bc60f Adjusted to new map update priorities
nipkow
parents: 67613
diff changeset
  1176
          apply (rule_tac xs = "(np a' x, h, (init_vars lvars)(pns[\<mapsto>]pvs, This\<mapsto>a'))" and st=blk in state_ok_exec)
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1177
             apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1178
            apply (simp (no_asm_simp) only: env_of_jmb_fst)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1179
           apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1180
          apply (simp (no_asm_use) only: env_of_jmb_fst)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1181
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1182
         (* show class_sig_defined G md (mn, pTs) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1183
         apply (simp (no_asm_simp) add: class_sig_defined_def)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1184
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1185
        (* show G,h \<turnstile> a' ::\<preceq> Class dynT *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1186
        apply (frule non_npD)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1187
         apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1188
        apply (erule exE)+
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1189
        apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1190
        apply (rule conf_obj_AddrI)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1191
        apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1192
        apply (rule widen_Class_Class [THEN iffD1], rule widen.refl)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1193
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1194
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1195
       (* show list_all2 (conf G h) pvs pTs *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1196
       apply (erule exE)+ apply (erule conjE)+
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1197
       apply (rule_tac Ts="pTsa" in conf_list_gext_widen)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1198
         apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1199
        apply (subgoal_tac "G \<turnstile> (gx s1, gs s1) -ps[\<succ>]pvs-> (x, h, l)")
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1200
         apply (frule_tac E="env_of_jmb G CL S" in evals_type_sound)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1201
             apply assumption+
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1202
            apply (simp only: env_of_jmb_fst)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1203
            apply (simp add: conforms_def xconf_def gs_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1204
           apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1205
          apply (simp (no_asm_use) only: gx_def gs_def surjective_pairing [symmetric])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1206
          apply (simp (no_asm_use) only: ty_exprs_list_all2)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1207
          apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1208
         apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1209
        apply (simp (no_asm_use) only: gx_def gs_def surjective_pairing [symmetric])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1210
       (* list_all2 (\<lambda>T T'. G \<turnstile> T \<preceq> T') pTsa pTs *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1211
       apply (rule max_spec_widen, simp only: env_of_jmb_fst)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1212
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1213
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1214
      (* show method (G, md) (mn, pTs) = Some (md, rT, pns, lvars, blk, res) *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1215
      apply (frule wf_prog_ws_prog [THEN method_in_md [THEN conjunct2]], assumption+)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1216
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1217
     (* show G\<turnstile>Class dynT \<preceq> Class md *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1218
     apply (simp (no_asm_use) only: widen_Class_Class)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1219
     apply (rule method_wf_mdecl [THEN conjunct1], assumption+)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1220
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1221
    (* is_class G md *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1222
    apply (rule wf_prog_ws_prog [THEN method_in_md [THEN conjunct1]], assumption+)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1223
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1224
   (* show is_class G dynT *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1225
   apply (frule non_npD)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1226
    apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1227
   apply (erule exE)+
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1228
   apply (erule conjE)+
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1229
   apply simp
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1230
   apply (rule subcls_is_class2)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1231
    apply assumption
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1232
   apply (frule expr_class_is_class [rotated])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1233
    apply (simp only: env_of_jmb_fst)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1234
    apply (rule wf_prog_ws_prog, assumption)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1235
   apply (simp only: env_of_jmb_fst)
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1236
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1237
  (* show G,h \<turnstile> a' ::\<preceq> Class C *)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1238
  apply (simp only: wtpd_exprs_def, erule exE)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1239
  apply (frule evals_preserves_conf)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1240
       apply (rule eval_conf, assumption+)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1241
       apply (rule env_of_jmb_fst, assumption+)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1242
   apply (rule env_of_jmb_fst)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1243
  apply (simp only: gh_conv)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1244
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1245
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1246
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1247
theorem compiler_correctness_eval: "
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1248
  \<lbrakk> G \<turnstile> (None,hp,loc) -ex \<succ> val-> (None,hp',loc');
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1249
  wf_java_prog G;
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1250
  class_sig_defined G C S;
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1251
  wtpd_expr (env_of_jmb G C S) ex;
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1252
  (None,hp,loc) ::\<preceq> (env_of_jmb G C S) \<rbrakk> \<Longrightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1253
  {(TranslComp.comp G), C, S} \<turnstile> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1254
    {hp, os, (locvars_locals G C S loc)}
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1255
      >- (compExpr (gmb G C S) ex) \<rightarrow> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1256
    {hp', val#os, (locvars_locals G C S loc')}"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1257
  apply (frule compiler_correctness [THEN conjunct1])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1258
  apply (auto simp: gh_def gx_def gs_def gl_def locvars_xstate_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1259
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1260
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1261
theorem compiler_correctness_exec: "
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 20432
diff changeset
  1262
  \<lbrakk> G \<turnstile> Norm (hp, loc) -st-> Norm (hp', loc');
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1263
  wf_java_prog G;
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1264
  class_sig_defined G C S;
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1265
  wtpd_stmt (env_of_jmb G C S) st;
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1266
  (None,hp,loc) ::\<preceq> (env_of_jmb G C S) \<rbrakk> \<Longrightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1267
  {(TranslComp.comp G), C, S} \<turnstile> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1268
    {hp, os, (locvars_locals G C S loc)}
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1269
      >- (compStmt (gmb G C S) st) \<rightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1270
    {hp', os, (locvars_locals G C S loc')}"
60304
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1271
  apply (frule compiler_correctness [THEN conjunct2 [THEN conjunct2]])
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1272
  apply (auto simp: gh_def gx_def gs_def gl_def locvars_xstate_def)
3f429b7d8eb5 modernized (slightly) type compiler in MicroJava
kleing
parents: 59199
diff changeset
  1273
  done
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1274
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1275
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1276
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1277
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1278
(* reinstall pair splits *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1279
declare split_paired_All [simp] split_paired_Ex [simp]
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1280
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
  1281
declare wf_prog_ws_prog [simp del]
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 14025
diff changeset
  1282
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
  1283
end