src/HOL/MicroJava/J/Conform.thy
author haftmann
Mon, 01 Mar 2010 13:40:23 +0100
changeset 35416 d8d7d1b785af
parent 30235 58d147683393
child 35417 47ee18b6ae32
permissions -rw-r--r--
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/J/Conform.thy
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     2
    Author:     David von Oheimb
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     3
    Copyright   1999 Technische Universitaet Muenchen
11070
cc421547e744 improved document (added headers etc)
oheimb
parents: 11026
diff changeset
     4
*)
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     5
12911
704713ca07ea new document
kleing
parents: 12888
diff changeset
     6
header {* \isaheader{Conformity Relations for Type Soundness Proof} *}
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     7
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14174
diff changeset
     8
theory Conform imports State WellType Exceptions begin
9346
297dcbf64526 re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents: 8082
diff changeset
     9
24783
5a3e336a2e37 avoid internal names;
wenzelm
parents: 16417
diff changeset
    10
types 'c env' = "'c prog \<times> (vname \<rightharpoonup> ty)"  -- "same as @{text env} of @{text WellType.thy}"
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    11
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 30235
diff changeset
    12
definition hext :: "aheap => aheap => bool" ("_ <=| _" [51,51] 50) where
11372
648795477bb5 corrected xsymbol/HTML syntax
oheimb
parents: 11070
diff changeset
    13
 "h<=|h' == \<forall>a C fs. h a = Some(C,fs) --> (\<exists>fs'. h' a = Some(C,fs'))"
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    14
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 30235
diff changeset
    15
definition conf :: "'c prog => aheap => val => ty => bool" 
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 30235
diff changeset
    16
                                   ("_,_ |- _ ::<= _"  [51,51,51,51] 50) where
30235
58d147683393 Made Option a separate theory and renamed option_map to Option.map
nipkow
parents: 26342
diff changeset
    17
 "G,h|-v::<=T == \<exists>T'. typeof (Option.map obj_ty o h) v = Some T' \<and> G\<turnstile>T'\<preceq>T"
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    18
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 30235
diff changeset
    19
definition lconf :: "'c prog => aheap => ('a \<rightharpoonup> val) => ('a \<rightharpoonup> ty) => bool"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 30235
diff changeset
    20
                                   ("_,_ |- _ [::<=] _" [51,51,51,51] 50) where
11372
648795477bb5 corrected xsymbol/HTML syntax
oheimb
parents: 11070
diff changeset
    21
 "G,h|-vs[::<=]Ts == \<forall>n T. Ts n = Some T --> (\<exists>v. vs n = Some v \<and> G,h|-v::<=T)"
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    22
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 30235
diff changeset
    23
definition oconf :: "'c prog => aheap => obj => bool" ("_,_ |- _ [ok]" [51,51,51] 50) where
11372
648795477bb5 corrected xsymbol/HTML syntax
oheimb
parents: 11070
diff changeset
    24
 "G,h|-obj [ok] == G,h|-snd obj[::<=]map_of (fields (G,fst obj))"
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    25
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 30235
diff changeset
    26
definition hconf :: "'c prog => aheap => bool" ("_ |-h _ [ok]" [51,51] 50) where
11372
648795477bb5 corrected xsymbol/HTML syntax
oheimb
parents: 11070
diff changeset
    27
 "G|-h h [ok]    == \<forall>a obj. h a = Some obj --> G,h|-obj [ok]"
13672
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
    28
 
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 30235
diff changeset
    29
definition xconf :: "aheap \<Rightarrow> val option \<Rightarrow> bool" where
13672
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
    30
  "xconf hp vo  == preallocated hp \<and> (\<forall> v. (vo = Some v) \<longrightarrow> (\<exists> xc. v = (Addr (XcptRef xc))))"
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    31
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 30235
diff changeset
    32
definition conforms :: "xstate => java_mb env' => bool" ("_ ::<= _" [51,51] 50) where
13672
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
    33
 "s::<=E == prg E|-h heap (store s) [ok] \<and> 
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
    34
            prg E,heap (store s)|-locals (store s)[::<=]localT E \<and> 
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
    35
            xconf (heap (store s)) (abrupt s)"
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    36
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    37
11372
648795477bb5 corrected xsymbol/HTML syntax
oheimb
parents: 11070
diff changeset
    38
syntax (xsymbols)
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    39
  hext     :: "aheap => aheap => bool"
11372
648795477bb5 corrected xsymbol/HTML syntax
oheimb
parents: 11070
diff changeset
    40
              ("_ \<le>| _" [51,51] 50)
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    41
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    42
  conf     :: "'c prog => aheap => val => ty => bool"
11372
648795477bb5 corrected xsymbol/HTML syntax
oheimb
parents: 11070
diff changeset
    43
              ("_,_ \<turnstile> _ ::\<preceq> _" [51,51,51,51] 50)
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    44
14134
0fdf5708c7a8 Replaced \<leadsto> by \<rightharpoonup>
nipkow
parents: 13672
diff changeset
    45
  lconf    :: "'c prog => aheap => ('a \<rightharpoonup> val) => ('a \<rightharpoonup> ty) => bool"
11372
648795477bb5 corrected xsymbol/HTML syntax
oheimb
parents: 11070
diff changeset
    46
              ("_,_ \<turnstile> _ [::\<preceq>] _" [51,51,51,51] 50)
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    47
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    48
  oconf    :: "'c prog => aheap => obj => bool"
11372
648795477bb5 corrected xsymbol/HTML syntax
oheimb
parents: 11070
diff changeset
    49
              ("_,_ \<turnstile> _ \<surd>" [51,51,51] 50)
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    50
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    51
  hconf    :: "'c prog => aheap => bool"
11372
648795477bb5 corrected xsymbol/HTML syntax
oheimb
parents: 11070
diff changeset
    52
              ("_ \<turnstile>h _ \<surd>" [51,51] 50)
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    53
24783
5a3e336a2e37 avoid internal names;
wenzelm
parents: 16417
diff changeset
    54
  conforms :: "state => java_mb env' => bool"
11372
648795477bb5 corrected xsymbol/HTML syntax
oheimb
parents: 11070
diff changeset
    55
              ("_ ::\<preceq> _" [51,51] 50)
10061
fe82134773dc added HTML syntax; added spaces in normal syntax for better documents
kleing
parents: 10042
diff changeset
    56
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    57
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    58
section "hext"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    59
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    60
lemma hextI: 
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    61
" \<forall>a C fs . h  a = Some (C,fs) -->   
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    62
      (\<exists>fs'. h' a = Some (C,fs')) ==> h\<le>|h'"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    63
apply (unfold hext_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    64
apply auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    65
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    66
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    67
lemma hext_objD: "[|h\<le>|h'; h a = Some (C,fs) |] ==> \<exists>fs'. h' a = Some (C,fs')"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    68
apply (unfold hext_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    69
apply (force)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    70
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    71
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    72
lemma hext_refl [simp]: "h\<le>|h"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    73
apply (rule hextI)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    74
apply (fast)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    75
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    76
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    77
lemma hext_new [simp]: "h a = None ==> h\<le>|h(a\<mapsto>x)"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    78
apply (rule hextI)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    79
apply auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    80
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    81
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    82
lemma hext_trans: "[|h\<le>|h'; h'\<le>|h''|] ==> h\<le>|h''"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    83
apply (rule hextI)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    84
apply (fast dest: hext_objD)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    85
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    86
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    87
lemma hext_upd_obj: "h a = Some (C,fs) ==> h\<le>|h(a\<mapsto>(C,fs'))"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    88
apply (rule hextI)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    89
apply auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    90
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    91
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    92
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    93
section "conf"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    94
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    95
lemma conf_Null [simp]: "G,h\<turnstile>Null::\<preceq>T = G\<turnstile>RefT NullT\<preceq>T"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    96
apply (unfold conf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    97
apply (simp (no_asm))
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    98
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
    99
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   100
lemma conf_litval [rule_format (no_asm), simp]: 
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   101
  "typeof (\<lambda>v. None) v = Some T --> G,h\<turnstile>v::\<preceq>T"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   102
apply (unfold conf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   103
apply (rule val.induct)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   104
apply auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   105
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   106
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   107
lemma conf_AddrI: "[|h a = Some obj; G\<turnstile>obj_ty obj\<preceq>T|] ==> G,h\<turnstile>Addr a::\<preceq>T"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   108
apply (unfold conf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   109
apply (simp)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   110
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   111
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   112
lemma conf_obj_AddrI: "[|h a = Some (C,fs); G\<turnstile>C\<preceq>C D|] ==> G,h\<turnstile>Addr a::\<preceq> Class D"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   113
apply (unfold conf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   114
apply (simp)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   115
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   116
12517
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   117
lemma defval_conf [rule_format (no_asm)]: 
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   118
  "is_type G T --> G,h\<turnstile>default_val T::\<preceq>T"
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   119
apply (unfold conf_def)
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 14134
diff changeset
   120
apply (rule_tac y = "T" in ty.exhaust)
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   121
apply  (erule ssubst)
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 14134
diff changeset
   122
apply  (rule_tac y = "prim_ty" in prim_ty.exhaust)
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   123
apply    (auto simp add: widen.null)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   124
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   125
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   126
lemma conf_upd_obj: 
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   127
"h a = Some (C,fs) ==> (G,h(a\<mapsto>(C,fs'))\<turnstile>x::\<preceq>T) = (G,h\<turnstile>x::\<preceq>T)"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   128
apply (unfold conf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   129
apply (rule val.induct)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   130
apply auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   131
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   132
12517
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   133
lemma conf_widen [rule_format (no_asm)]: 
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   134
  "wf_prog wf_mb G ==> G,h\<turnstile>x::\<preceq>T --> G\<turnstile>T\<preceq>T' --> G,h\<turnstile>x::\<preceq>T'"
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   135
apply (unfold conf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   136
apply (rule val.induct)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   137
apply (auto intro: widen_trans)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   138
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   139
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   140
lemma conf_hext [rule_format (no_asm)]: "h\<le>|h' ==> G,h\<turnstile>v::\<preceq>T --> G,h'\<turnstile>v::\<preceq>T"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   141
apply (unfold conf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   142
apply (rule val.induct)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   143
apply (auto dest: hext_objD)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   144
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   145
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   146
lemma new_locD: "[|h a = None; G,h\<turnstile>Addr t::\<preceq>T|] ==> t\<noteq>a"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   147
apply (unfold conf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   148
apply auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   149
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   150
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   151
lemma conf_RefTD [rule_format (no_asm)]: 
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   152
 "G,h\<turnstile>a'::\<preceq>RefT T --> a' = Null |   
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   153
  (\<exists>a obj T'. a' = Addr a \<and>  h a = Some obj \<and>  obj_ty obj = T' \<and>  G\<turnstile>T'\<preceq>RefT T)"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   154
apply (unfold conf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   155
apply(induct_tac "a'")
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   156
apply(auto)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   157
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   158
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   159
lemma conf_NullTD: "G,h\<turnstile>a'::\<preceq>RefT NullT ==> a' = Null"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   160
apply (drule conf_RefTD)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   161
apply auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   162
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   163
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   164
lemma non_npD: "[|a' \<noteq> Null; G,h\<turnstile>a'::\<preceq>RefT t|] ==>  
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   165
  \<exists>a C fs. a' = Addr a \<and>  h a = Some (C,fs) \<and>  G\<turnstile>Class C\<preceq>RefT t"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   166
apply (drule conf_RefTD)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   167
apply auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   168
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   169
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
   170
lemma non_np_objD: "!!G. [|a' \<noteq> Null; G,h\<turnstile>a'::\<preceq> Class C|] ==>  
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   171
  (\<exists>a C' fs. a' = Addr a \<and>  h a = Some (C',fs) \<and>  G\<turnstile>C'\<preceq>C C)"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   172
apply (fast dest: non_npD)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   173
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   174
12517
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   175
lemma non_np_objD' [rule_format (no_asm)]: 
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   176
  "a' \<noteq> Null ==> wf_prog wf_mb G ==> G,h\<turnstile>a'::\<preceq>RefT t --> 
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   177
  (\<exists>a C fs. a' = Addr a \<and>  h a = Some (C,fs) \<and>  G\<turnstile>Class C\<preceq>RefT t)"
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 14134
diff changeset
   178
apply(rule_tac y = "t" in ref_ty.exhaust)
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   179
 apply (fast dest: conf_NullTD)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   180
apply (fast dest: non_np_objD)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   181
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   182
12517
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   183
lemma conf_list_gext_widen [rule_format (no_asm)]: 
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   184
  "wf_prog wf_mb G ==> \<forall>Ts Ts'. list_all2 (conf G h) vs Ts --> 
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   185
  list_all2 (\<lambda>T T'. G\<turnstile>T\<preceq>T') Ts Ts' -->  list_all2 (conf G h) vs Ts'"
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   186
apply(induct_tac "vs")
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   187
 apply(clarsimp)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   188
apply(clarsimp)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   189
apply(frule list_all2_lengthD [THEN sym])
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   190
apply(simp (no_asm_use) add: length_Suc_conv)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   191
apply(safe)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   192
apply(frule list_all2_lengthD [THEN sym])
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   193
apply(simp (no_asm_use) add: length_Suc_conv)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   194
apply(clarify)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   195
apply(fast elim: conf_widen)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   196
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   197
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   198
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   199
section "lconf"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   200
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   201
lemma lconfD: "[| G,h\<turnstile>vs[::\<preceq>]Ts; Ts n = Some T |] ==> G,h\<turnstile>(the (vs n))::\<preceq>T"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   202
apply (unfold lconf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   203
apply (force)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   204
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   205
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   206
lemma lconf_hext [elim]: "[| G,h\<turnstile>l[::\<preceq>]L; h\<le>|h' |] ==> G,h'\<turnstile>l[::\<preceq>]L"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   207
apply (unfold lconf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   208
apply  (fast elim: conf_hext)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   209
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   210
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   211
lemma lconf_upd: "!!X. [| G,h\<turnstile>l[::\<preceq>]lT;  
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   212
  G,h\<turnstile>v::\<preceq>T; lT va = Some T |] ==> G,h\<turnstile>l(va\<mapsto>v)[::\<preceq>]lT"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   213
apply (unfold lconf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   214
apply auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   215
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   216
12517
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   217
lemma lconf_init_vars_lemma [rule_format (no_asm)]: 
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   218
  "\<forall>x. P x --> R (dv x) x ==> (\<forall>x. map_of fs f = Some x --> P x) -->  
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   219
  (\<forall>T. map_of fs f = Some T -->  
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   220
  (\<exists>v. map_of (map (\<lambda>(f,ft). (f, dv ft)) fs) f = Some v \<and>  R v T))"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   221
apply( induct_tac "fs")
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   222
apply auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   223
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   224
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   225
lemma lconf_init_vars [intro!]: 
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   226
"\<forall>n. \<forall>T. map_of fs n = Some T --> is_type G T ==> G,h\<turnstile>init_vars fs[::\<preceq>]map_of fs"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   227
apply (unfold lconf_def init_vars_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   228
apply auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   229
apply( rule lconf_init_vars_lemma)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   230
apply(   erule_tac [3] asm_rl)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   231
apply(  intro strip)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   232
apply(  erule defval_conf)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   233
apply auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   234
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   235
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   236
lemma lconf_ext: "[|G,s\<turnstile>l[::\<preceq>]L; G,s\<turnstile>v::\<preceq>T|] ==> G,s\<turnstile>l(vn\<mapsto>v)[::\<preceq>]L(vn\<mapsto>T)"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   237
apply (unfold lconf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   238
apply auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   239
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   240
12517
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   241
lemma lconf_ext_list [rule_format (no_asm)]: 
12888
f6c1e7306c40 nodups -> distinct
nipkow
parents: 12517
diff changeset
   242
  "G,h\<turnstile>l[::\<preceq>]L ==> \<forall>vs Ts. distinct vns --> length Ts = length vns --> 
12517
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   243
  list_all2 (\<lambda>v T. G,h\<turnstile>v::\<preceq>T) vs Ts --> G,h\<turnstile>l(vns[\<mapsto>]vs)[::\<preceq>]L(vns[\<mapsto>]Ts)"
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   244
apply (unfold lconf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   245
apply( induct_tac "vns")
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   246
apply(  clarsimp)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   247
apply( clarsimp)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   248
apply( frule list_all2_lengthD)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   249
apply( auto simp add: length_Suc_conv)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   250
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   251
13672
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   252
lemma lconf_restr: "\<lbrakk>lT vn = None; G, h \<turnstile> l [::\<preceq>] lT(vn\<mapsto>T)\<rbrakk> \<Longrightarrow> G, h \<turnstile> l [::\<preceq>] lT"
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   253
apply (unfold lconf_def)
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   254
apply (intro strip)
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   255
apply (case_tac "n = vn")
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   256
apply auto
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   257
done
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   258
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   259
section "oconf"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   260
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   261
lemma oconf_hext: "G,h\<turnstile>obj\<surd> ==> h\<le>|h' ==> G,h'\<turnstile>obj\<surd>"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   262
apply (unfold oconf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   263
apply (fast)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   264
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   265
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   266
lemma oconf_obj: "G,h\<turnstile>(C,fs)\<surd> =  
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   267
  (\<forall>T f. map_of(fields (G,C)) f = Some T --> (\<exists>v. fs f = Some v \<and>  G,h\<turnstile>v::\<preceq>T))"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   268
apply (unfold oconf_def lconf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   269
apply auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   270
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   271
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   272
lemmas oconf_objD = oconf_obj [THEN iffD1, THEN spec, THEN spec, THEN mp]
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   273
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   274
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   275
section "hconf"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   276
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   277
lemma hconfD: "[|G\<turnstile>h h\<surd>; h a = Some obj|] ==> G,h\<turnstile>obj\<surd>"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   278
apply (unfold hconf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   279
apply (fast)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   280
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   281
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   282
lemma hconfI: "\<forall>a obj. h a=Some obj --> G,h\<turnstile>obj\<surd> ==> G\<turnstile>h h\<surd>"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   283
apply (unfold hconf_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   284
apply (fast)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   285
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   286
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   287
13672
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   288
section "xconf"
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   289
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   290
lemma xconf_raise_if: "xconf h x \<Longrightarrow> xconf h (raise_if b xcn x)"
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   291
by (simp add: xconf_def raise_if_def)
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   292
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   293
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   294
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   295
section "conforms"
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   296
13672
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   297
lemma conforms_heapD: "(x, (h, l))::\<preceq>(G, lT) ==> G\<turnstile>h h\<surd>"
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   298
apply (unfold conforms_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   299
apply (simp)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   300
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   301
13672
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   302
lemma conforms_localD: "(x, (h, l))::\<preceq>(G, lT) ==> G,h\<turnstile>l[::\<preceq>]lT"
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   303
apply (unfold conforms_def)
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   304
apply (simp)
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   305
done
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   306
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   307
lemma conforms_xcptD: "(x, (h, l))::\<preceq>(G, lT) ==> xconf h x"
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   308
apply (unfold conforms_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   309
apply (simp)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   310
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   311
13672
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   312
lemma conformsI: "[|G\<turnstile>h h\<surd>; G,h\<turnstile>l[::\<preceq>]lT; xconf h x|] ==> (x, (h, l))::\<preceq>(G, lT)"
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   313
apply (unfold conforms_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   314
apply auto
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   315
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   316
13672
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   317
lemma conforms_restr: "\<lbrakk>lT vn = None; s ::\<preceq> (G, lT(vn\<mapsto>T)) \<rbrakk> \<Longrightarrow> s ::\<preceq> (G, lT)"
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   318
by (simp add: conforms_def, fast intro: lconf_restr)
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   319
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   320
lemma conforms_xcpt_change: "\<lbrakk> (x, (h,l))::\<preceq> (G, lT); xconf h x \<longrightarrow> xconf h x' \<rbrakk> \<Longrightarrow> (x', (h,l))::\<preceq> (G, lT)"
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   321
by (simp add: conforms_def)
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   322
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   323
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   324
lemma preallocated_hext: "\<lbrakk> preallocated h; h\<le>|h'\<rbrakk> \<Longrightarrow> preallocated h'"
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   325
by (simp add: preallocated_def hext_def)
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   326
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   327
lemma xconf_hext: "\<lbrakk> xconf h vo; h\<le>|h'\<rbrakk> \<Longrightarrow> xconf h' vo"
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   328
by (simp add: xconf_def preallocated_def hext_def)
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   329
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   330
lemma conforms_hext: "[|(x,(h,l))::\<preceq>(G,lT); h\<le>|h'; G\<turnstile>h h'\<surd> |] 
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   331
  ==> (x,(h',l))::\<preceq>(G,lT)"
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   332
by( fast dest: conforms_localD conforms_xcptD elim!: conformsI xconf_hext)
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   333
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   334
12517
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   335
lemma conforms_upd_obj: 
13672
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   336
  "[|(x,(h,l))::\<preceq>(G, lT); G,h(a\<mapsto>obj)\<turnstile>obj\<surd>; h\<le>|h(a\<mapsto>obj)|] 
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   337
  ==> (x,(h(a\<mapsto>obj),l))::\<preceq>(G, lT)"
12517
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   338
apply(rule conforms_hext)
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   339
apply  auto
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   340
apply(rule hconfI)
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   341
apply(drule conforms_heapD)
360e3215f029 exception merge, cleanup, tuned
kleing
parents: 11372
diff changeset
   342
apply(tactic {* auto_tac (HOL_cs addEs [thm "oconf_hext"] 
26342
0f65fa163304 more antiquotations;
wenzelm
parents: 24783
diff changeset
   343
                addDs [thm "hconfD"], @{simpset} delsimps [split_paired_All]) *})
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   344
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   345
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   346
lemma conforms_upd_local: 
13672
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   347
"[|(x,(h, l))::\<preceq>(G, lT); G,h\<turnstile>v::\<preceq>T; lT va = Some T|] 
b95d12325b51 Added compiler
streckem
parents: 12951
diff changeset
   348
  ==> (x,(h, l(va\<mapsto>v)))::\<preceq>(G, lT)"
11026
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   349
apply (unfold conforms_def)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   350
apply( auto elim: lconf_upd)
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   351
done
a50365d21144 converted to Isar, simplifying recursion on class hierarchy
oheimb
parents: 10061
diff changeset
   352
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
   353
end