src/HOL/MicroJava/BV/Correct.thy
author kleing
Thu, 21 Sep 2000 10:42:49 +0200
changeset 10042 7164dc0d24d8
parent 9941 fe05af7ec816
child 10056 9f84ffa4a8d0
permissions -rw-r--r--
unsymbolized
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/BV/Correct.thy
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     2
    ID:         $Id$
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     3
    Author:     Cornelia Pusch
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     4
    Copyright   1999 Technische Universitaet Muenchen
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     5
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     6
The invariant for the type safety proof.
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     7
*)
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     8
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
     9
header "Type Safety Invariant"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    10
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    11
theory Correct = BVSpec:
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    12
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    13
constdefs
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    14
 approx_val :: "[jvm_prog,aheap,val,ty err] => bool"
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    15
"approx_val G h v any == case any of Err => True | Ok T => G,h\<turnstile>v::\<preceq>T"
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    16
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    17
 approx_loc :: "[jvm_prog,aheap,val list,locvars_type] => bool"
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    18
"approx_loc G hp loc LT == list_all2 (approx_val G hp) loc LT"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    19
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    20
 approx_stk :: "[jvm_prog,aheap,opstack,opstack_type] => bool"
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    21
"approx_stk G hp stk ST == approx_loc G hp stk (map Ok ST)"
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    22
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    23
 correct_frame  :: "[jvm_prog,aheap,state_type,nat,bytecode] => frame => bool"
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    24
"correct_frame G hp == \<lambda>(ST,LT) maxl ins (stk,loc,C,sig,pc).
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    25
   approx_stk G hp stk ST  \<and> approx_loc G hp loc LT \<and> 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    26
   pc < length ins \<and> length loc=length(snd sig)+maxl+1"
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    27
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    28
 correct_frame_opt :: "[jvm_prog,aheap,state_type option,nat,bytecode] => frame => bool"
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    29
"correct_frame_opt G hp s == case s of None => \<lambda>maxl ins f. False | Some t => correct_frame G hp t"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    30
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    31
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    32
consts
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    33
 correct_frames  :: "[jvm_prog,aheap,prog_type,ty,sig,frame list] => bool"
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    34
primrec
8045
816f566c414f Fixed a problem with returning from the last frame.
nipkow
parents: 8034
diff changeset
    35
"correct_frames G hp phi rT0 sig0 [] = True"
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    36
8045
816f566c414f Fixed a problem with returning from the last frame.
nipkow
parents: 8034
diff changeset
    37
"correct_frames G hp phi rT0 sig0 (f#frs) =
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    38
	(let (stk,loc,C,sig,pc) = f in
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    39
  (\<exists>ST LT rT maxl ins.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    40
    phi C sig ! pc = Some (ST,LT) \<and> 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    41
    method (G,C) sig = Some(C,rT,(maxl,ins)) \<and>
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    42
	(\<exists>C' mn pTs k. pc = k+1 \<and> ins!k = (Invoke C' mn pTs) \<and>
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    43
         (mn,pTs) = sig0 \<and> 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    44
         (\<exists>apTs D ST' LT'.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    45
         (phi C sig)!k = Some ((rev apTs) @ (Class D) # ST', LT') \<and>
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    46
         length apTs = length pTs \<and>
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    47
         (\<exists>D' rT' maxl' ins'.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    48
           method (G,D) sig0 = Some(D',rT',(maxl',ins')) \<and>
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    49
           G \<turnstile> rT0 \<preceq> rT') \<and>
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    50
	 correct_frame G hp (tl ST, LT) maxl ins f \<and> 
8045
816f566c414f Fixed a problem with returning from the last frame.
nipkow
parents: 8034
diff changeset
    51
	 correct_frames G hp phi rT sig frs))))"
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    52
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    53
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    54
constdefs
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    55
 correct_state :: "[jvm_prog,prog_type,jvm_state] => bool"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    56
                  ("_,_\<turnstile>JVM _\<surd>"  [51,51] 50)
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    57
"correct_state G phi == \<lambda>(xp,hp,frs).
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    58
   case xp of
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    59
     None => (case frs of
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    60
	           [] => True
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    61
             | (f#fs) => G\<turnstile>h hp\<surd> \<and>
8045
816f566c414f Fixed a problem with returning from the last frame.
nipkow
parents: 8034
diff changeset
    62
			(let (stk,loc,C,sig,pc) = f
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    63
		         in
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    64
                         \<exists>rT maxl ins s.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    65
                         method (G,C) sig = Some(C,rT,(maxl,ins)) \<and>
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    66
                         phi C sig ! pc = Some s \<and>
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    67
			 correct_frame G hp s maxl ins f \<and> 
8045
816f566c414f Fixed a problem with returning from the last frame.
nipkow
parents: 8034
diff changeset
    68
		         correct_frames G hp phi rT sig fs))
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    69
   | Some x => True" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    70
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    71
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    72
lemma sup_heap_newref:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    73
  "hp x = None ==> hp \<le>| hp(newref hp \<mapsto> obj)"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    74
apply (unfold hext_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    75
apply clarsimp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    76
apply (drule newref_None 1) back
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    77
apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    78
.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    79
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    80
lemma sup_heap_update_value:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    81
  "hp a = Some (C,od') ==> hp \<le>| hp (a \<mapsto> (C,od))"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    82
by (simp add: hext_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    83
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    84
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    85
(** approx_val **)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    86
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    87
lemma approx_val_Err:
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    88
  "approx_val G hp x Err"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    89
by (simp add: approx_val_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    90
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    91
lemma approx_val_Null:
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    92
  "approx_val G hp Null (Ok (RefT x))"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    93
by (auto intro: null simp add: approx_val_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    94
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
    95
lemma approx_val_imp_approx_val_assConvertible [rule_format]: 
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
    96
  "wf_prog wt G ==> approx_val G hp v (Ok T) --> G\<turnstile> T\<preceq>T' --> approx_val G hp v (Ok T')"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    97
by (cases T) (auto intro: conf_widen simp add: approx_val_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
    98
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
    99
lemma approx_val_imp_approx_val_sup_heap [rule_format]:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   100
  "approx_val G hp v at --> hp \<le>| hp' --> approx_val G hp' v at"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   101
apply (simp add: approx_val_def split: err.split)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   102
apply (blast intro: conf_hext)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   103
.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   104
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   105
lemma approx_val_imp_approx_val_heap_update:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   106
  "[|hp a = Some obj'; G,hp\<turnstile> v::\<preceq>T; obj_ty obj = obj_ty obj'|] 
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   107
  ==> G,hp(a\<mapsto>obj)\<turnstile> v::\<preceq>T"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   108
by (cases v, auto simp add: obj_ty_def conf_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   109
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   110
lemma approx_val_imp_approx_val_sup [rule_format]:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   111
  "wf_prog wt G ==> (approx_val G h v us) --> (G \<turnstile> us <=o us') --> (approx_val G h v us')"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   112
apply (simp add: sup_PTS_eq approx_val_def split: err.split)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   113
apply (blast intro: conf_widen)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   114
.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   115
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   116
lemma approx_loc_imp_approx_val_sup:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   117
  "[|wf_prog wt G; approx_loc G hp loc LT; idx < length LT; v = loc!idx; G \<turnstile> LT!idx <=o at|]
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   118
  ==> approx_val G hp v at"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   119
apply (unfold approx_loc_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   120
apply (unfold list_all2_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   121
apply (auto intro: approx_val_imp_approx_val_sup simp add: split_def all_set_conv_all_nth)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   122
.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   123
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   124
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   125
(** approx_loc **)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   126
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   127
lemma approx_loc_Cons [iff]:
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   128
  "approx_loc G hp (s#xs) (l#ls) = (approx_val G hp s l \<and> approx_loc G hp xs ls)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   129
by (simp add: approx_loc_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   130
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   131
lemma assConv_approx_stk_imp_approx_loc [rule_format]:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   132
  "wf_prog wt G ==> (\<forall>tt'\<in>set (zip tys_n ts). tt' \<in> widen G) 
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   133
  --> length tys_n = length ts --> approx_stk G hp s tys_n --> 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   134
  approx_loc G hp s (map Ok ts)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   135
apply (unfold approx_stk_def approx_loc_def list_all2_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   136
apply (clarsimp simp add: all_set_conv_all_nth)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   137
apply (rule approx_val_imp_approx_val_assConvertible)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   138
apply auto
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   139
.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   140
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   141
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   142
lemma approx_loc_imp_approx_loc_sup_heap [rule_format]:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   143
  "\<forall>lvars. approx_loc G hp lvars lt --> hp \<le>| hp' --> approx_loc G hp' lvars lt"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   144
apply (unfold approx_loc_def list_all2_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   145
apply (cases lt)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   146
 apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   147
apply clarsimp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   148
apply (rule approx_val_imp_approx_val_sup_heap)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   149
apply auto
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   150
.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   151
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   152
lemma approx_loc_imp_approx_loc_sup [rule_format]:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   153
  "wf_prog wt G ==> approx_loc G hp lvars lt --> G \<turnstile> lt <=l lt' --> approx_loc G hp lvars lt'"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   154
apply (unfold sup_loc_def approx_loc_def list_all2_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   155
apply (auto simp add: all_set_conv_all_nth)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   156
apply (auto elim: approx_val_imp_approx_val_sup)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   157
.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   158
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   159
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   160
lemma approx_loc_imp_approx_loc_subst [rule_format]:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   161
  "\<forall>loc idx x X. (approx_loc G hp loc LT) --> (approx_val G hp x X) 
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   162
  --> (approx_loc G hp (loc[idx:=x]) (LT[idx:=X]))"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   163
apply (unfold approx_loc_def list_all2_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   164
apply (auto dest: subsetD [OF set_update_subset_insert] simp add: zip_update)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   165
.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   166
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   167
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   168
lemmas [cong] = conj_cong 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   169
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   170
lemma approx_loc_append [rule_format]:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   171
  "\<forall>L1 l2 L2. length l1=length L1 --> 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   172
  approx_loc G hp (l1@l2) (L1@L2) = (approx_loc G hp l1 L1 \<and> approx_loc G hp l2 L2)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   173
apply (unfold approx_loc_def list_all2_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   174
apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   175
apply blast
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   176
.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   177
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   178
lemmas [cong del] = conj_cong
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   179
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   180
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   181
(** approx_stk **)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   182
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   183
lemma approx_stk_rev_lem:
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   184
  "approx_stk G hp (rev s) (rev t) = approx_stk G hp s t"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   185
apply (unfold approx_stk_def approx_loc_def list_all2_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   186
apply (auto simp add: zip_rev sym [OF rev_map])
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   187
.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   188
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   189
lemma approx_stk_rev:
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   190
  "approx_stk G hp (rev s) t = approx_stk G hp s (rev t)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   191
by (auto intro: subst [OF approx_stk_rev_lem])
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   192
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   193
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   194
lemma approx_stk_imp_approx_stk_sup_heap [rule_format]:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   195
  "\<forall>lvars. approx_stk G hp lvars lt --> hp \<le>| hp' --> approx_stk G hp' lvars lt"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   196
by (auto intro: approx_loc_imp_approx_loc_sup_heap simp add: approx_stk_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   197
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   198
lemma approx_stk_imp_approx_stk_sup [rule_format]:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   199
  "wf_prog wt G ==> approx_stk G hp lvars st --> (G \<turnstile> map Ok st <=l (map Ok st')) 
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   200
  --> approx_stk G hp lvars st'" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   201
by (auto intro: approx_loc_imp_approx_loc_sup simp add: approx_stk_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   202
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   203
lemma approx_stk_Nil [iff]:
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   204
  "approx_stk G hp [] []"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   205
by (simp add: approx_stk_def approx_loc_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   206
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   207
lemma approx_stk_Cons [iff]:
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   208
  "approx_stk G hp (x # stk) (S#ST) = 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   209
   (approx_val G hp x (Ok S) \<and> approx_stk G hp stk ST)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   210
by (simp add: approx_stk_def approx_loc_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   211
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   212
lemma approx_stk_Cons_lemma [iff]:
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   213
  "approx_stk G hp stk (S#ST') = 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   214
  (\<exists>s stk'. stk = s#stk' \<and> approx_val G hp s (Ok S) \<and> approx_stk G hp stk' ST')"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   215
by (simp add: list_all2_Cons2 approx_stk_def approx_loc_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   216
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   217
lemma approx_stk_append_lemma:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   218
  "approx_stk G hp stk (S@ST') ==>
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   219
   (\<exists>s stk'. stk = s@stk' \<and> length s = length S \<and> length stk' = length ST' \<and> 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   220
             approx_stk G hp s S \<and> approx_stk G hp stk' ST')"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   221
by (simp add: list_all2_append2 approx_stk_def approx_loc_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   222
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   223
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   224
(** oconf **)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   225
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   226
lemma correct_init_obj:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   227
  "[|is_class G C; wf_prog wt G|] ==> 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   228
  G,h \<turnstile> (C, map_of (map (\<lambda>(f,fT).(f,default_val fT)) (fields(G,C)))) \<surd>"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   229
apply (unfold oconf_def lconf_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   230
apply (simp add: map_of_map)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   231
apply (force intro: defval_conf dest: map_of_SomeD is_type_fields)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   232
.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   233
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   234
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   235
lemma oconf_imp_oconf_field_update [rule_format]:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   236
  "[|map_of (fields (G, oT)) FD = Some T; G,hp\<turnstile>v::\<preceq>T; G,hp\<turnstile>(oT,fs)\<surd> |]
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   237
  ==> G,hp\<turnstile>(oT, fs(FD\<mapsto>v))\<surd>"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   238
by (simp add: oconf_def lconf_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   239
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   240
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   241
lemma oconf_imp_oconf_heap_newref [rule_format]:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   242
"hp x = None --> G,hp\<turnstile>obj\<surd> --> G,hp\<turnstile>obj'\<surd> --> G,(hp(newref hp\<mapsto>obj'))\<turnstile>obj\<surd>"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   243
apply (unfold oconf_def lconf_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   244
apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   245
apply (fast intro: conf_hext sup_heap_newref)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   246
.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   247
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   248
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   249
lemma oconf_imp_oconf_heap_update [rule_format]:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   250
  "hp a = Some obj' --> obj_ty obj' = obj_ty obj'' --> G,hp\<turnstile>obj\<surd> 
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   251
  --> G,hp(a\<mapsto>obj'')\<turnstile>obj\<surd>"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   252
apply (unfold oconf_def lconf_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   253
apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   254
apply (force intro: approx_val_imp_approx_val_heap_update)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   255
.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   256
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   257
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   258
(** hconf **)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   259
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   260
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   261
lemma hconf_imp_hconf_newref [rule_format]:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   262
  "hp x = None --> G\<turnstile>h hp\<surd> --> G,hp\<turnstile>obj\<surd> --> G\<turnstile>h hp(newref hp\<mapsto>obj)\<surd>"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   263
apply (simp add: hconf_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   264
apply (fast intro: oconf_imp_oconf_heap_newref)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   265
.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   266
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   267
lemma hconf_imp_hconf_field_update [rule_format]:
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   268
  "map_of (fields (G, oT)) (F, D) = Some T \<and> hp oloc = Some(oT,fs) \<and> 
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   269
  G,hp\<turnstile>v::\<preceq>T \<and> G\<turnstile>h hp\<surd> --> G\<turnstile>h hp(oloc \<mapsto> (oT, fs((F,D)\<mapsto>v)))\<surd>"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   270
apply (simp add: hconf_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   271
apply (force intro: oconf_imp_oconf_heap_update oconf_imp_oconf_field_update 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   272
             simp add: obj_ty_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   273
.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   274
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   275
(** correct_frames **)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   276
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   277
lemmas [simp del] = fun_upd_apply
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   278
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   279
lemma correct_frames_imp_correct_frames_field_update [rule_format]:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   280
  "\<forall>rT C sig. correct_frames G hp phi rT sig frs --> 
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   281
  hp a = Some (C,od) --> map_of (fields (G, C)) fl = Some fd --> 
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   282
  G,hp\<turnstile>v::\<preceq>fd 
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   283
  --> correct_frames G (hp(a \<mapsto> (C, od(fl\<mapsto>v)))) phi rT sig frs";
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   284
apply (induct frs)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   285
 apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   286
apply (clarsimp simp add: correct_frame_def) (*takes long*)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   287
apply (intro exI conjI)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   288
     apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   289
    apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   290
   apply force
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   291
  apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   292
 apply (rule approx_stk_imp_approx_stk_sup_heap)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   293
  apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   294
 apply (rule sup_heap_update_value)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   295
 apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   296
apply (rule approx_loc_imp_approx_loc_sup_heap)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   297
 apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   298
apply (rule sup_heap_update_value)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   299
apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   300
.
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   301
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   302
lemma correct_frames_imp_correct_frames_newref [rule_format]:
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   303
  "\<forall>rT C sig. hp x = None --> correct_frames G hp phi rT sig frs \<and> oconf G hp obj 
7164dc0d24d8 unsymbolized
kleing
parents: 9941
diff changeset
   304
  --> correct_frames G (hp(newref hp \<mapsto> obj)) phi rT sig frs"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   305
apply (induct frs)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   306
 apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   307
apply (clarsimp simp add: correct_frame_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   308
apply (intro exI conjI)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   309
     apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   310
    apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   311
   apply force
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   312
  apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   313
 apply (rule approx_stk_imp_approx_stk_sup_heap)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   314
  apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   315
 apply (rule sup_heap_newref)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   316
 apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   317
apply (rule approx_loc_imp_approx_loc_sup_heap)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   318
 apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   319
apply (rule sup_heap_newref)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   320
apply simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   321
.
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
   322
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
   323
end
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9549
diff changeset
   324