src/HOL/MicroJava/BV/Effect.thy
author Fabian Huch <huch@in.tum.de>
Mon, 13 Nov 2023 17:00:13 +0100
changeset 78972 7a39f151e9a7
parent 69597 ff784d5a5bfb
permissions -rw-r--r--
proper parallel paths for timing heuristic;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/BV/Effect.thy
d09d0f160888 exceptions
kleing
parents:
diff changeset
     2
    Author:     Gerwin Klein
d09d0f160888 exceptions
kleing
parents:
diff changeset
     3
    Copyright   2000 Technische Universitaet Muenchen
d09d0f160888 exceptions
kleing
parents:
diff changeset
     4
*)
d09d0f160888 exceptions
kleing
parents:
diff changeset
     5
61361
8b5f00202e1a isabelle update_cartouches;
wenzelm
parents: 58886
diff changeset
     6
section \<open>Effect of Instructions on the State Type\<close>
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
     7
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13717
diff changeset
     8
theory Effect 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13717
diff changeset
     9
imports JVMType "../JVM/JVMExceptions"
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13717
diff changeset
    10
begin
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13717
diff changeset
    11
42463
f270e3e18be5 modernized specifications;
wenzelm
parents: 35440
diff changeset
    12
type_synonym succ_type = "(p_count \<times> state_type option) list"
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
    13
61361
8b5f00202e1a isabelle update_cartouches;
wenzelm
parents: 58886
diff changeset
    14
text \<open>Program counter of successor instructions:\<close>
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    15
primrec succs :: "instr \<Rightarrow> p_count \<Rightarrow> p_count list" where
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
    16
  "succs (Load idx) pc         = [pc+1]"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    17
| "succs (Store idx) pc        = [pc+1]"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    18
| "succs (LitPush v) pc        = [pc+1]"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    19
| "succs (Getfield F C) pc     = [pc+1]"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    20
| "succs (Putfield F C) pc     = [pc+1]"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    21
| "succs (New C) pc            = [pc+1]"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    22
| "succs (Checkcast C) pc      = [pc+1]"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    23
| "succs Pop pc                = [pc+1]"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    24
| "succs Dup pc                = [pc+1]"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    25
| "succs Dup_x1 pc             = [pc+1]"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    26
| "succs Dup_x2 pc             = [pc+1]"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    27
| "succs Swap pc               = [pc+1]"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    28
| "succs IAdd pc               = [pc+1]"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    29
| "succs (Ifcmpeq b) pc        = [pc+1, nat (int pc + b)]"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    30
| "succs (Goto b) pc           = [nat (int pc + b)]"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    31
| "succs Return pc             = [pc]"  
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    32
| "succs (Invoke C mn fpTs) pc = [pc+1]"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    33
| "succs Throw pc              = [pc]"
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
    34
d09d0f160888 exceptions
kleing
parents:
diff changeset
    35
text "Effect of instruction on the state type:"
d09d0f160888 exceptions
kleing
parents:
diff changeset
    36
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    37
fun eff' :: "instr \<times> jvm_prog \<times> state_type \<Rightarrow> state_type"
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    38
where
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    39
"eff' (Load idx,  G, (ST, LT))          = (ok_val (LT ! idx) # ST, LT)" |
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    40
"eff' (Store idx, G, (ts#ST, LT))       = (ST, LT[idx:= OK ts])" |
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    41
"eff' (LitPush v, G, (ST, LT))           = (the (typeof (\<lambda>v. None) v) # ST, LT)" |
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    42
"eff' (Getfield F C, G, (oT#ST, LT))    = (snd (the (field (G,C) F)) # ST, LT)" |
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    43
"eff' (Putfield F C, G, (vT#oT#ST, LT)) = (ST,LT)" |
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    44
"eff' (New C, G, (ST,LT))               = (Class C # ST, LT)" |
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    45
"eff' (Checkcast C, G, (RefT rt#ST,LT)) = (Class C # ST,LT)" |
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    46
"eff' (Pop, G, (ts#ST,LT))              = (ST,LT)" |
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    47
"eff' (Dup, G, (ts#ST,LT))              = (ts#ts#ST,LT)" |
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    48
"eff' (Dup_x1, G, (ts1#ts2#ST,LT))      = (ts1#ts2#ts1#ST,LT)" |
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    49
"eff' (Dup_x2, G, (ts1#ts2#ts3#ST,LT))  = (ts1#ts2#ts3#ts1#ST,LT)" |
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    50
"eff' (Swap, G, (ts1#ts2#ST,LT))        = (ts2#ts1#ST,LT)" |
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
    51
"eff' (IAdd, G, (PrimT Integer#PrimT Integer#ST,LT)) 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    52
                                         = (PrimT Integer#ST,LT)" |
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    53
"eff' (Ifcmpeq b, G, (ts1#ts2#ST,LT))   = (ST,LT)" |
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    54
"eff' (Goto b, G, s)                    = s" |
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62390
diff changeset
    55
  \<comment> \<open>Return has no successor instruction in the same method\<close>
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    56
"eff' (Return, G, s)                    = s" |
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 62390
diff changeset
    57
  \<comment> \<open>Throw always terminates abruptly\<close>
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    58
"eff' (Throw, G, s)                     = s" |
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
    59
"eff' (Invoke C mn fpTs, G, (ST,LT))    = (let ST' = drop (length fpTs) ST 
d09d0f160888 exceptions
kleing
parents:
diff changeset
    60
  in  (fst (snd (the (method (G,C) (mn,fpTs))))#(tl ST'),LT))" 
d09d0f160888 exceptions
kleing
parents:
diff changeset
    61
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    62
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    63
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    64
primrec match_any :: "jvm_prog \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> cname list" where
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
    65
  "match_any G pc [] = []"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    66
| "match_any G pc (e#es) = (let (start_pc, end_pc, handler_pc, catch_type) = e;
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
    67
                                es' = match_any G pc es 
d09d0f160888 exceptions
kleing
parents:
diff changeset
    68
                            in 
d09d0f160888 exceptions
kleing
parents:
diff changeset
    69
                            if start_pc <= pc \<and> pc < end_pc then catch_type#es' else es')"
d09d0f160888 exceptions
kleing
parents:
diff changeset
    70
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    71
primrec match :: "jvm_prog \<Rightarrow> xcpt \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> cname list" where
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    72
  "match G X pc [] = []"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    73
| "match G X pc (e#es) = 
13717
78f91fcdf560 beautified "match"
kleing
parents: 13066
diff changeset
    74
  (if match_exception_entry G (Xcpt X) pc e then [Xcpt X] else match G X pc es)"
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    75
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    76
lemma match_some_entry:
13717
78f91fcdf560 beautified "match"
kleing
parents: 13066
diff changeset
    77
  "match G X pc et = (if \<exists>e \<in> set et. match_exception_entry G (Xcpt X) pc e then [Xcpt X] else [])"
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    78
  by (induct et) auto
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
    79
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    80
fun
13006
51c5f3f11d16 symbolized
kleing
parents: 12974
diff changeset
    81
  xcpt_names :: "instr \<times> jvm_prog \<times> p_count \<times> exception_table \<Rightarrow> cname list" 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    82
where
13717
78f91fcdf560 beautified "match"
kleing
parents: 13066
diff changeset
    83
  "xcpt_names (Getfield F C, G, pc, et) = match G NullPointer pc et" 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    84
| "xcpt_names (Putfield F C, G, pc, et) = match G NullPointer pc et" 
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    85
| "xcpt_names (New C, G, pc, et)        = match G OutOfMemory pc et"
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    86
| "xcpt_names (Checkcast C, G, pc, et)  = match G ClassCast pc et"
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    87
| "xcpt_names (Throw, G, pc, et)        = match_any G pc et"
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    88
| "xcpt_names (Invoke C m p, G, pc, et) = match_any G pc et" 
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
    89
| "xcpt_names (i, G, pc, et)            = []" 
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
    90
d09d0f160888 exceptions
kleing
parents:
diff changeset
    91
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    92
definition xcpt_eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> p_count \<Rightarrow> state_type option \<Rightarrow> exception_table \<Rightarrow> succ_type" where
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
    93
  "xcpt_eff i G pc s et == 
d09d0f160888 exceptions
kleing
parents:
diff changeset
    94
   map (\<lambda>C. (the (match_exception_table G C pc et), case s of None \<Rightarrow> None | Some s' \<Rightarrow> Some ([Class C], snd s'))) 
d09d0f160888 exceptions
kleing
parents:
diff changeset
    95
       (xcpt_names (i,G,pc,et))"
d09d0f160888 exceptions
kleing
parents:
diff changeset
    96
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    97
definition norm_eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> state_type option \<Rightarrow> state_type option" where
55466
786edc984c98 merged 'Option.map' and 'Option.map_option'
blanchet
parents: 54863
diff changeset
    98
  "norm_eff i G == map_option (\<lambda>s. eff' (i,G,s))"
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
    99
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
   100
definition eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> state_type option \<Rightarrow> succ_type" where
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   101
  "eff i G pc et s == (map (\<lambda>pc'. (pc',norm_eff i G s)) (succs i pc)) @ (xcpt_eff i G pc s et)"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   102
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
   103
definition isPrimT :: "ty \<Rightarrow> bool" where
12772
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   104
  "isPrimT T == case T of PrimT T' \<Rightarrow> True | RefT T' \<Rightarrow> False"
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   105
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
   106
definition isRefT :: "ty \<Rightarrow> bool" where
12772
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   107
  "isRefT T == case T of PrimT T' \<Rightarrow> False | RefT T' \<Rightarrow> True"
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   108
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   109
lemma isPrimT [simp]:
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   110
  "isPrimT T = (\<exists>T'. T = PrimT T')" by (simp add: isPrimT_def split: ty.splits)
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   111
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   112
lemma isRefT [simp]:
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   113
  "isRefT T = (\<exists>T'. T = RefT T')" by (simp add: isRefT_def split: ty.splits)
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   114
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   115
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   116
lemma "list_all2 P a b \<Longrightarrow> \<forall>(x,y) \<in> set (zip a b). P x y"
55524
f41ef840f09d folded 'list_all2' with the relator generated by 'datatype_new'
blanchet
parents: 55466
diff changeset
   117
  by (simp add: list_all2_iff) 
12772
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   118
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   119
d09d0f160888 exceptions
kleing
parents:
diff changeset
   120
text "Conditions under which eff is applicable:"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   121
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   122
fun
13006
51c5f3f11d16 symbolized
kleing
parents: 12974
diff changeset
   123
app' :: "instr \<times> jvm_prog \<times> p_count \<times> nat \<times> ty \<times> state_type \<Rightarrow> bool"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   124
where
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   125
"app' (Load idx, G, pc, maxs, rT, s) = 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   126
  (idx < length (snd s) \<and> (snd s) ! idx \<noteq> Err \<and> length (fst s) < maxs)" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   127
"app' (Store idx, G, pc, maxs, rT, (ts#ST, LT)) = 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   128
  (idx < length LT)" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   129
"app' (LitPush v, G, pc, maxs, rT, s) = 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   130
  (length (fst s) < maxs \<and> typeof (\<lambda>t. None) v \<noteq> None)" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   131
"app' (Getfield F C, G, pc, maxs, rT, (oT#ST, LT)) = 
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   132
  (is_class G C \<and> field (G,C) F \<noteq> None \<and> fst (the (field (G,C) F)) = C \<and> 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   133
  G \<turnstile> oT \<preceq> (Class C))" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   134
"app' (Putfield F C, G, pc, maxs, rT, (vT#oT#ST, LT)) = 
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   135
  (is_class G C \<and> field (G,C) F \<noteq> None \<and> fst (the (field (G,C) F)) = C \<and>
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   136
  G \<turnstile> oT \<preceq> (Class C) \<and> G \<turnstile> vT \<preceq> (snd (the (field (G,C) F))))" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   137
"app' (New C, G, pc, maxs, rT, s) = 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   138
  (is_class G C \<and> length (fst s) < maxs)" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   139
"app' (Checkcast C, G, pc, maxs, rT, (RefT rt#ST,LT)) = 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   140
  (is_class G C)" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   141
"app' (Pop, G, pc, maxs, rT, (ts#ST,LT)) = 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   142
  True" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   143
"app' (Dup, G, pc, maxs, rT, (ts#ST,LT)) = 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   144
  (1+length ST < maxs)" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   145
"app' (Dup_x1, G, pc, maxs, rT, (ts1#ts2#ST,LT)) = 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   146
  (2+length ST < maxs)" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   147
"app' (Dup_x2, G, pc, maxs, rT, (ts1#ts2#ts3#ST,LT)) = 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   148
  (3+length ST < maxs)" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   149
"app' (Swap, G, pc, maxs, rT, (ts1#ts2#ST,LT)) = 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   150
  True" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   151
"app' (IAdd, G, pc, maxs, rT, (PrimT Integer#PrimT Integer#ST,LT)) =
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   152
  True" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   153
"app' (Ifcmpeq b, G, pc, maxs, rT, (ts#ts'#ST,LT)) = 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   154
  (0 \<le> int pc + b \<and> (isPrimT ts \<and> ts' = ts \<or> isRefT ts \<and> isRefT ts'))" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   155
"app' (Goto b, G, pc, maxs, rT, s) = 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   156
  (0 \<le> int pc + b)" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   157
"app' (Return, G, pc, maxs, rT, (T#ST,LT)) = 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   158
  (G \<turnstile> T \<preceq> rT)" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   159
"app' (Throw, G, pc, maxs, rT, (T#ST,LT)) = 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   160
  isRefT T" |
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   161
"app' (Invoke C mn fpTs, G, pc, maxs, rT, s) = 
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   162
  (length fpTs < length (fst s) \<and> 
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   163
  (let apTs = rev (take (length fpTs) (fst s));
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   164
       X    = hd (drop (length fpTs) (fst s)) 
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   165
   in  
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   166
       G \<turnstile> X \<preceq> Class C \<and> is_class G C \<and> method (G,C) (mn,fpTs) \<noteq> None \<and>
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   167
       list_all2 (\<lambda>x y. G \<turnstile> x \<preceq> y) apTs fpTs))" |
12772
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   168
  
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   169
"app' (i,G, pc,maxs,rT,s) = False"
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   170
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
   171
definition xcpt_app :: "instr \<Rightarrow> jvm_prog \<Rightarrow> nat \<Rightarrow> exception_table \<Rightarrow> bool" where
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   172
  "xcpt_app i G pc et \<equiv> \<forall>C\<in>set(xcpt_names (i,G,pc,et)). is_class G C"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   173
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
   174
definition app :: "instr \<Rightarrow> jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> exception_table \<Rightarrow> state_type option \<Rightarrow> bool" where
13006
51c5f3f11d16 symbolized
kleing
parents: 12974
diff changeset
   175
  "app i G maxs rT pc et s == case s of None \<Rightarrow> True | Some t \<Rightarrow> app' (i,G,pc,maxs,rT,t) \<and> xcpt_app i G pc et"
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   176
d09d0f160888 exceptions
kleing
parents:
diff changeset
   177
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   178
lemma match_any_match_table:
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   179
  "C \<in> set (match_any G pc et) \<Longrightarrow> match_exception_table G C pc et \<noteq> None"
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   180
  apply (induct et)
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   181
   apply simp
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   182
  apply simp
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   183
  apply clarify
62390
842917225d56 more canonical names
nipkow
parents: 62042
diff changeset
   184
  apply (simp split: if_split_asm)
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   185
  apply (auto simp add: match_exception_entry_def)
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   186
  done
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   187
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   188
lemma match_X_match_table:
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   189
  "C \<in> set (match G X pc et) \<Longrightarrow> match_exception_table G C pc et \<noteq> None"
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   190
  apply (induct et)
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   191
   apply simp
62390
842917225d56 more canonical names
nipkow
parents: 62042
diff changeset
   192
  apply (simp split: if_split_asm)
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   193
  done
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   194
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   195
lemma xcpt_names_in_et:
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   196
  "C \<in> set (xcpt_names (i,G,pc,et)) \<Longrightarrow> 
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   197
  \<exists>e \<in> set et. the (match_exception_table G C pc et) = fst (snd (snd e))"
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   198
  apply (cases i)
18576
8d98b7711e47 Reversed Larry's option/iff change.
nipkow
parents: 18447
diff changeset
   199
  apply (auto dest!: match_any_match_table match_X_match_table 
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   200
              dest: match_exception_table_in_et)
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   201
  done
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   202
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   203
13006
51c5f3f11d16 symbolized
kleing
parents: 12974
diff changeset
   204
lemma 1: "2 < length a \<Longrightarrow> (\<exists>l l' l'' ls. a = l#l'#l''#ls)"
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   205
proof (cases a)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   206
  fix x xs assume "a = x#xs" "2 < length a"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   207
  thus ?thesis by - (cases xs, simp, cases "tl xs", auto)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   208
qed auto
d09d0f160888 exceptions
kleing
parents:
diff changeset
   209
13006
51c5f3f11d16 symbolized
kleing
parents: 12974
diff changeset
   210
lemma 2: "\<not>(2 < length a) \<Longrightarrow> a = [] \<or> (\<exists> l. a = [l]) \<or> (\<exists> l l'. a = [l,l'])"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   211
proof -
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   212
  assume "\<not>(2 < length a)"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   213
  hence "length a < (Suc (Suc (Suc 0)))" by simp
d09d0f160888 exceptions
kleing
parents:
diff changeset
   214
  hence * : "length a = 0 \<or> length a = Suc 0 \<or> length a = Suc (Suc 0)" 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   215
    by (auto simp add: less_Suc_eq)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   216
d09d0f160888 exceptions
kleing
parents:
diff changeset
   217
  { 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   218
    fix x 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   219
    assume "length x = Suc 0"
46720
9722171731af tuned proofs;
wenzelm
parents: 42463
diff changeset
   220
    hence "\<exists> l. x = [l]"  by (cases x) auto
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   221
  } note 0 = this
d09d0f160888 exceptions
kleing
parents:
diff changeset
   222
46720
9722171731af tuned proofs;
wenzelm
parents: 42463
diff changeset
   223
  have "length a = Suc (Suc 0) \<Longrightarrow> \<exists>l l'. a = [l,l']" by (cases a) (auto dest: 0)
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   224
  with * show ?thesis by (auto dest: 0)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   225
qed
d09d0f160888 exceptions
kleing
parents:
diff changeset
   226
d09d0f160888 exceptions
kleing
parents:
diff changeset
   227
lemmas [simp] = app_def xcpt_app_def
d09d0f160888 exceptions
kleing
parents:
diff changeset
   228
61361
8b5f00202e1a isabelle update_cartouches;
wenzelm
parents: 58886
diff changeset
   229
text \<open>
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   230
\medskip
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   231
simp rules for \<^term>\<open>app\<close>
61361
8b5f00202e1a isabelle update_cartouches;
wenzelm
parents: 58886
diff changeset
   232
\<close>
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   233
lemma appNone[simp]: "app i G maxs rT pc et None = True" by simp
d09d0f160888 exceptions
kleing
parents:
diff changeset
   234
d09d0f160888 exceptions
kleing
parents:
diff changeset
   235
d09d0f160888 exceptions
kleing
parents:
diff changeset
   236
lemma appLoad[simp]:
d09d0f160888 exceptions
kleing
parents:
diff changeset
   237
"(app (Load idx) G maxs rT pc et (Some s)) = (\<exists>ST LT. s = (ST,LT) \<and> idx < length LT \<and> LT!idx \<noteq> Err \<and> length ST < maxs)"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   238
  by (cases s, simp)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   239
d09d0f160888 exceptions
kleing
parents:
diff changeset
   240
lemma appStore[simp]:
d09d0f160888 exceptions
kleing
parents:
diff changeset
   241
"(app (Store idx) G maxs rT pc et (Some s)) = (\<exists>ts ST LT. s = (ts#ST,LT) \<and> idx < length LT)"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   242
  by (cases s, cases "2 < length (fst s)", auto dest: 1 2)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   243
d09d0f160888 exceptions
kleing
parents:
diff changeset
   244
lemma appLitPush[simp]:
d09d0f160888 exceptions
kleing
parents:
diff changeset
   245
"(app (LitPush v) G maxs rT pc et (Some s)) = (\<exists>ST LT. s = (ST,LT) \<and> length ST < maxs \<and> typeof (\<lambda>v. None) v \<noteq> None)"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   246
  by (cases s, simp)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   247
d09d0f160888 exceptions
kleing
parents:
diff changeset
   248
lemma appGetField[simp]:
d09d0f160888 exceptions
kleing
parents:
diff changeset
   249
"(app (Getfield F C) G maxs rT pc et (Some s)) = 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   250
 (\<exists> oT vT ST LT. s = (oT#ST, LT) \<and> is_class G C \<and>  
13717
78f91fcdf560 beautified "match"
kleing
parents: 13066
diff changeset
   251
  field (G,C) F = Some (C,vT) \<and> G \<turnstile> oT \<preceq> (Class C) \<and> (\<forall>x \<in> set (match G NullPointer pc et). is_class G x))"
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   252
  by (cases s, cases "2 <length (fst s)", auto dest!: 1 2)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   253
d09d0f160888 exceptions
kleing
parents:
diff changeset
   254
lemma appPutField[simp]:
d09d0f160888 exceptions
kleing
parents:
diff changeset
   255
"(app (Putfield F C) G maxs rT pc et (Some s)) = 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   256
 (\<exists> vT vT' oT ST LT. s = (vT#oT#ST, LT) \<and> is_class G C \<and> 
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
   257
  field (G,C) F = Some (C, vT') \<and> G \<turnstile> oT \<preceq> (Class C) \<and> G \<turnstile> vT \<preceq> vT' \<and>
13717
78f91fcdf560 beautified "match"
kleing
parents: 13066
diff changeset
   258
  (\<forall>x \<in> set (match G NullPointer pc et). is_class G x))"
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   259
  by (cases s, cases "2 <length (fst s)", auto dest!: 1 2)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   260
d09d0f160888 exceptions
kleing
parents:
diff changeset
   261
lemma appNew[simp]:
d09d0f160888 exceptions
kleing
parents:
diff changeset
   262
  "(app (New C) G maxs rT pc et (Some s)) = 
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
   263
  (\<exists>ST LT. s=(ST,LT) \<and> is_class G C \<and> length ST < maxs \<and>
13717
78f91fcdf560 beautified "match"
kleing
parents: 13066
diff changeset
   264
  (\<forall>x \<in> set (match G OutOfMemory pc et). is_class G x))"
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   265
  by (cases s, simp)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   266
d09d0f160888 exceptions
kleing
parents:
diff changeset
   267
lemma appCheckcast[simp]: 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   268
  "(app (Checkcast C) G maxs rT pc et (Some s)) =  
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
   269
  (\<exists>rT ST LT. s = (RefT rT#ST,LT) \<and> is_class G C \<and>
13717
78f91fcdf560 beautified "match"
kleing
parents: 13066
diff changeset
   270
  (\<forall>x \<in> set (match G ClassCast pc et). is_class G x))"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   271
  by (cases s, cases "fst s", simp) (cases "hd (fst s)", auto)
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   272
d09d0f160888 exceptions
kleing
parents:
diff changeset
   273
lemma appPop[simp]: 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   274
"(app Pop G maxs rT pc et (Some s)) = (\<exists>ts ST LT. s = (ts#ST,LT))"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   275
  by (cases s, cases "2 <length (fst s)", auto dest: 1 2)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   276
d09d0f160888 exceptions
kleing
parents:
diff changeset
   277
d09d0f160888 exceptions
kleing
parents:
diff changeset
   278
lemma appDup[simp]:
d09d0f160888 exceptions
kleing
parents:
diff changeset
   279
"(app Dup G maxs rT pc et (Some s)) = (\<exists>ts ST LT. s = (ts#ST,LT) \<and> 1+length ST < maxs)" 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   280
  by (cases s, cases "2 <length (fst s)", auto dest: 1 2)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   281
d09d0f160888 exceptions
kleing
parents:
diff changeset
   282
d09d0f160888 exceptions
kleing
parents:
diff changeset
   283
lemma appDup_x1[simp]:
d09d0f160888 exceptions
kleing
parents:
diff changeset
   284
"(app Dup_x1 G maxs rT pc et (Some s)) = (\<exists>ts1 ts2 ST LT. s = (ts1#ts2#ST,LT) \<and> 2+length ST < maxs)" 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   285
  by (cases s, cases "2 <length (fst s)", auto dest: 1 2)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   286
d09d0f160888 exceptions
kleing
parents:
diff changeset
   287
d09d0f160888 exceptions
kleing
parents:
diff changeset
   288
lemma appDup_x2[simp]:
d09d0f160888 exceptions
kleing
parents:
diff changeset
   289
"(app Dup_x2 G maxs rT pc et (Some s)) = (\<exists>ts1 ts2 ts3 ST LT. s = (ts1#ts2#ts3#ST,LT) \<and> 3+length ST < maxs)"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   290
  by (cases s, cases "2 <length (fst s)", auto dest: 1 2)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   291
d09d0f160888 exceptions
kleing
parents:
diff changeset
   292
d09d0f160888 exceptions
kleing
parents:
diff changeset
   293
lemma appSwap[simp]:
d09d0f160888 exceptions
kleing
parents:
diff changeset
   294
"app Swap G maxs rT pc et (Some s) = (\<exists>ts1 ts2 ST LT. s = (ts1#ts2#ST,LT))" 
25362
8d06e11a01d1 tuned proofs -- avoid open cases;
wenzelm
parents: 18576
diff changeset
   295
  by (cases s, cases "2 <length (fst s)") (auto dest: 1 2)
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   296
d09d0f160888 exceptions
kleing
parents:
diff changeset
   297
d09d0f160888 exceptions
kleing
parents:
diff changeset
   298
lemma appIAdd[simp]:
d09d0f160888 exceptions
kleing
parents:
diff changeset
   299
"app IAdd G maxs rT pc et (Some s) = (\<exists> ST LT. s = (PrimT Integer#PrimT Integer#ST,LT))"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   300
  (is "?app s = ?P s")
25362
8d06e11a01d1 tuned proofs -- avoid open cases;
wenzelm
parents: 18576
diff changeset
   301
proof (cases s)
8d06e11a01d1 tuned proofs -- avoid open cases;
wenzelm
parents: 18576
diff changeset
   302
  case (Pair a b)
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   303
  have "?app (a,b) = ?P (a,b)"
25362
8d06e11a01d1 tuned proofs -- avoid open cases;
wenzelm
parents: 18576
diff changeset
   304
  proof (cases a)
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   305
    fix t ts assume a: "a = t#ts"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   306
    show ?thesis
d09d0f160888 exceptions
kleing
parents:
diff changeset
   307
    proof (cases t)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   308
      fix p assume p: "t = PrimT p"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   309
      show ?thesis
d09d0f160888 exceptions
kleing
parents:
diff changeset
   310
      proof (cases p)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   311
        assume ip: "p = Integer"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   312
        show ?thesis
d09d0f160888 exceptions
kleing
parents:
diff changeset
   313
        proof (cases ts)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   314
          fix t' ts' assume t': "ts = t' # ts'"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   315
          show ?thesis
d09d0f160888 exceptions
kleing
parents:
diff changeset
   316
          proof (cases t')
d09d0f160888 exceptions
kleing
parents:
diff changeset
   317
            fix p' assume "t' = PrimT p'"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   318
            with t' ip p a
25362
8d06e11a01d1 tuned proofs -- avoid open cases;
wenzelm
parents: 18576
diff changeset
   319
            show ?thesis by (cases p') auto
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   320
          qed (auto simp add: a p ip t')
d09d0f160888 exceptions
kleing
parents:
diff changeset
   321
        qed (auto simp add: a p ip)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   322
      qed (auto simp add: a p)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   323
    qed (auto simp add: a)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   324
  qed auto
d09d0f160888 exceptions
kleing
parents:
diff changeset
   325
  with Pair show ?thesis by simp
d09d0f160888 exceptions
kleing
parents:
diff changeset
   326
qed
d09d0f160888 exceptions
kleing
parents:
diff changeset
   327
d09d0f160888 exceptions
kleing
parents:
diff changeset
   328
d09d0f160888 exceptions
kleing
parents:
diff changeset
   329
lemma appIfcmpeq[simp]:
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   330
"app (Ifcmpeq b) G maxs rT pc et (Some s) = 
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   331
  (\<exists>ts1 ts2 ST LT. s = (ts1#ts2#ST,LT) \<and> 0 \<le> int pc + b \<and>
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   332
  ((\<exists> p. ts1 = PrimT p \<and> ts2 = PrimT p) \<or> (\<exists>r r'. ts1 = RefT r \<and> ts2 = RefT r')))" 
12772
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   333
  by (cases s, cases "2 <length (fst s)", auto dest!: 1 2)
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   334
d09d0f160888 exceptions
kleing
parents:
diff changeset
   335
lemma appReturn[simp]:
d09d0f160888 exceptions
kleing
parents:
diff changeset
   336
"app Return G maxs rT pc et (Some s) = (\<exists>T ST LT. s = (T#ST,LT) \<and> (G \<turnstile> T \<preceq> rT))" 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   337
  by (cases s, cases "2 <length (fst s)", auto dest: 1 2)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   338
d09d0f160888 exceptions
kleing
parents:
diff changeset
   339
lemma appGoto[simp]:
12974
7acfab8361d1 enforce positive branch targets
kleing
parents: 12951
diff changeset
   340
"app (Goto b) G maxs rT pc et (Some s) = (0 \<le> int pc + b)"
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   341
  by simp
d09d0f160888 exceptions
kleing
parents:
diff changeset
   342
d09d0f160888 exceptions
kleing
parents:
diff changeset
   343
lemma appThrow[simp]:
d09d0f160888 exceptions
kleing
parents:
diff changeset
   344
  "app Throw G maxs rT pc et (Some s) = 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   345
  (\<exists>T ST LT r. s=(T#ST,LT) \<and> T = RefT r \<and> (\<forall>C \<in> set (match_any G pc et). is_class G C))"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   346
  by (cases s, cases "2 < length (fst s)", auto dest: 1 2)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   347
d09d0f160888 exceptions
kleing
parents:
diff changeset
   348
lemma appInvoke[simp]:
d09d0f160888 exceptions
kleing
parents:
diff changeset
   349
"app (Invoke C mn fpTs) G maxs rT pc et (Some s) = (\<exists>apTs X ST LT mD' rT' b'.
d09d0f160888 exceptions
kleing
parents:
diff changeset
   350
  s = ((rev apTs) @ (X # ST), LT) \<and> length apTs = length fpTs \<and> is_class G C \<and>
d09d0f160888 exceptions
kleing
parents:
diff changeset
   351
  G \<turnstile> X \<preceq> Class C \<and> (\<forall>(aT,fT)\<in>set(zip apTs fpTs). G \<turnstile> aT \<preceq> fT) \<and> 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   352
  method (G,C) (mn,fpTs) = Some (mD', rT', b') \<and> 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   353
  (\<forall>C \<in> set (match_any G pc et). is_class G C))" (is "?app s = ?P s")
25362
8d06e11a01d1 tuned proofs -- avoid open cases;
wenzelm
parents: 18576
diff changeset
   354
proof (cases s)
55524
f41ef840f09d folded 'list_all2' with the relator generated by 'datatype_new'
blanchet
parents: 55466
diff changeset
   355
  note list_all2_iff [simp]
25362
8d06e11a01d1 tuned proofs -- avoid open cases;
wenzelm
parents: 18576
diff changeset
   356
  case (Pair a b)
13006
51c5f3f11d16 symbolized
kleing
parents: 12974
diff changeset
   357
  have "?app (a,b) \<Longrightarrow> ?P (a,b)"
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   358
  proof -
d09d0f160888 exceptions
kleing
parents:
diff changeset
   359
    assume app: "?app (a,b)"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   360
    hence "a = (rev (rev (take (length fpTs) a))) @ (drop (length fpTs) a) \<and> 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   361
           length fpTs < length a" (is "?a \<and> ?l") 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   362
      by auto
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   363
    hence "?a \<and> 0 < length (drop (length fpTs) a)" (is "?a \<and> ?l") 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   364
      by auto
d09d0f160888 exceptions
kleing
parents:
diff changeset
   365
    hence "?a \<and> ?l \<and> length (rev (take (length fpTs) a)) = length fpTs" 
32443
16464c3f86bd tuned proofs
nipkow
parents: 30235
diff changeset
   366
      by (auto)
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   367
    hence "\<exists>apTs ST. a = rev apTs @ ST \<and> length apTs = length fpTs \<and> 0 < length ST" 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   368
      by blast
d09d0f160888 exceptions
kleing
parents:
diff changeset
   369
    hence "\<exists>apTs ST. a = rev apTs @ ST \<and> length apTs = length fpTs \<and> ST \<noteq> []" 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   370
      by blast
d09d0f160888 exceptions
kleing
parents:
diff changeset
   371
    hence "\<exists>apTs ST. a = rev apTs @ ST \<and> length apTs = length fpTs \<and> 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   372
           (\<exists>X ST'. ST = X#ST')" 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   373
      by (simp add: neq_Nil_conv)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   374
    hence "\<exists>apTs X ST. a = rev apTs @ X # ST \<and> length apTs = length fpTs" 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   375
      by blast
d09d0f160888 exceptions
kleing
parents:
diff changeset
   376
    with app
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   377
    show ?thesis by clarsimp blast
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   378
  qed
d09d0f160888 exceptions
kleing
parents:
diff changeset
   379
  with Pair 
12772
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   380
  have "?app s \<Longrightarrow> ?P s" by (simp only:)
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   381
  moreover
54863
82acc20ded73 prefer more canonical names for lemmas on min/max
haftmann
parents: 46720
diff changeset
   382
  have "?P s \<Longrightarrow> ?app s" by (clarsimp simp add: min.absorb2)
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   383
  ultimately
12772
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   384
  show ?thesis by (rule iffI) 
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   385
qed 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   386
d09d0f160888 exceptions
kleing
parents:
diff changeset
   387
lemma effNone: 
d09d0f160888 exceptions
kleing
parents:
diff changeset
   388
  "(pc', s') \<in> set (eff i G pc et None) \<Longrightarrow> s' = None"
d09d0f160888 exceptions
kleing
parents:
diff changeset
   389
  by (auto simp add: eff_def xcpt_eff_def norm_eff_def)
d09d0f160888 exceptions
kleing
parents:
diff changeset
   390
12772
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   391
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   392
lemma xcpt_app_lemma [code]:
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   393
  "xcpt_app i G pc et = list_all (is_class G) (xcpt_names (i, G, pc, et))"
17088
3f797ec16f02 list_all_conv -> iff
nipkow
parents: 15481
diff changeset
   394
  by (simp add: list_all_iff)
12772
7b7051ae49a0 tuned for directly executable definitions
kleing
parents: 12516
diff changeset
   395
12516
d09d0f160888 exceptions
kleing
parents:
diff changeset
   396
lemmas [simp del] = app_def xcpt_app_def
d09d0f160888 exceptions
kleing
parents:
diff changeset
   397
d09d0f160888 exceptions
kleing
parents:
diff changeset
   398
end