src/HOL/MicroJava/BV/Step.thy
author kleing
Wed, 09 Aug 2000 11:53:00 +0200
changeset 9559 1f99296758c2
parent 9549 40d64cb4f4e6
child 9580 d955914193e0
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/BV/Step.thy
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
     2
    ID:         $Id$
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
     3
    Author:     Gerwin Klein
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
     4
    Copyright   2000 Technische Universitaet Muenchen
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
     5
*)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
     6
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
     7
header {* Effect of instructions on the state type *}
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
     8
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
     9
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    10
theory Step = Convert :
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    11
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    12
9559
kleing
parents: 9549
diff changeset
    13
text "Effect of instruction on the state type"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    14
consts 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    15
step :: "instr \\<times> jvm_prog \\<times> state_type \\<Rightarrow> state_type option"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    16
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    17
recdef step "{}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    18
"step (Load idx,  G, (ST, LT))          = Some (the (LT ! idx) # ST, LT)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    19
"step (Store idx, G, (ts#ST, LT))       = Some (ST, LT[idx:= Some ts])"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    20
"step (Bipush i, G, (ST, LT))           = Some (PrimT Integer # ST, LT)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    21
"step (Aconst_null, G, (ST, LT))        = Some (NT#ST,LT)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    22
"step (Getfield F C, G, (oT#ST, LT))    = Some (snd (the (field (G,C) F)) # ST, LT)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    23
"step (Putfield F C, G, (vT#oT#ST, LT)) = Some (ST,LT)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    24
"step (New C, G, (ST,LT))               = Some (Class C # ST, LT)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    25
"step (Checkcast C, G, (RefT rt#ST,LT)) = Some (Class C # ST,LT)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    26
"step (Pop, G, (ts#ST,LT))              = Some (ST,LT)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    27
"step (Dup, G, (ts#ST,LT))              = Some (ts#ts#ST,LT)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    28
"step (Dup_x1, G, (ts1#ts2#ST,LT))      = Some (ts1#ts2#ts1#ST,LT)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    29
"step (Dup_x2, G, (ts1#ts2#ts3#ST,LT))  = Some (ts1#ts2#ts3#ts1#ST,LT)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    30
"step (Swap, G, (ts1#ts2#ST,LT))        = Some (ts2#ts1#ST,LT)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    31
"step (IAdd, G, (PrimT Integer#PrimT Integer#ST,LT)) 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    32
                                          = Some (PrimT Integer#ST,LT)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    33
"step (Ifcmpeq b, G, (ts1#ts2#ST,LT))   = Some (ST,LT)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    34
"step (Goto b, G, s)                    = Some s"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    35
"step (Return, G, (T#ST,LT))            = None"   (* Return has no successor instruction in the same method *)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    36
"step (Invoke C mn fpTs, G, (ST,LT))    = (let ST' = drop (length fpTs) ST in
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    37
                                              Some (fst (snd (the (method (G,C) (mn,fpTs))))#(tl ST'),LT))" 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    38
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    39
"step (i,G,s)                           = None"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    40
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    41
9559
kleing
parents: 9549
diff changeset
    42
text "Conditions under which step is applicable"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    43
consts
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    44
app :: "instr \\<times> jvm_prog \\<times> ty \\<times> state_type \\<Rightarrow> bool"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    45
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    46
recdef app "{}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    47
"app (Load idx, G, rT, s)                  = (idx < length (snd s) \\<and> (snd s) ! idx \\<noteq> None)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    48
"app (Store idx, G, rT, (ts#ST, LT))       = (idx < length LT)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    49
"app (Bipush i, G, rT, s)                  = True"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    50
"app (Aconst_null, G, rT, s)               = True"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    51
"app (Getfield F C, G, rT, (oT#ST, LT))    = (is_class G C \\<and> 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    52
                                              field (G,C) F \\<noteq> None \\<and>
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    53
                                              fst (the (field (G,C) F)) = C \\<and>
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    54
                                              G \\<turnstile> oT \\<preceq> (Class C))"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    55
"app (Putfield F C, G, rT, (vT#oT#ST, LT)) = (is_class G C \\<and> 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    56
                                              field (G,C) F \\<noteq> None \\<and>
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    57
                                              fst (the (field (G,C) F)) = C \\<and>
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    58
                                              G \\<turnstile> oT \\<preceq> (Class C) \\<and>
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    59
                                              G \\<turnstile> vT \\<preceq> (snd (the (field (G,C) F))))" 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    60
"app (New C, G, rT, s)                     = (is_class G C)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    61
"app (Checkcast C, G, rT, (RefT rt#ST,LT)) = (is_class G C)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    62
"app (Pop, G, rT, (ts#ST,LT))              = True"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    63
"app (Dup, G, rT, (ts#ST,LT))              = True"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    64
"app (Dup_x1, G, rT, (ts1#ts2#ST,LT))      = True"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    65
"app (Dup_x2, G, rT, (ts1#ts2#ts3#ST,LT))  = True"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    66
"app (Swap, G, rT, (ts1#ts2#ST,LT))        = True"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    67
"app (IAdd, G, rT, (PrimT Integer#PrimT Integer#ST,LT)) 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    68
                                           = True"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    69
"app (Ifcmpeq b, G, rT, (ts1#ts2#ST,LT))   = ((\\<exists> p. ts1 = PrimT p \\<and> ts1 = PrimT p) \\<or> 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    70
                                              (\\<exists>r r'. ts1 = RefT r \\<and> ts2 = RefT r'))"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    71
"app (Goto b, G, rT, s)                    = True"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    72
"app (Return, G, rT, (T#ST,LT))            = (G \\<turnstile> T \\<preceq> rT)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    73
app_invoke:
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    74
"app (Invoke C mn fpTs, G, rT, s)          = (length fpTs < length (fst s) \\<and> 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    75
                                              (let 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    76
                                                apTs = rev (take (length fpTs) (fst s));
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    77
                                                X    = hd (drop (length fpTs) (fst s)) 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    78
                                              in
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    79
                                                G \\<turnstile> X \\<preceq> Class C \\<and> 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    80
                                                (\\<forall>(aT,fT)\\<in>set(zip apTs fpTs). G \\<turnstile> aT \\<preceq> fT) \\<and>
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    81
                                                method (G,C) (mn,fpTs) \\<noteq> None
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    82
                                             ))"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    83
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    84
"app (i,G,rT,s)                            = False"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    85
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    86
9559
kleing
parents: 9549
diff changeset
    87
text {* \isa{p_count} of successor instructions *}
kleing
parents: 9549
diff changeset
    88
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    89
consts
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    90
succs :: "instr \\<Rightarrow> p_count \\<Rightarrow> p_count set"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    91
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    92
primrec 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    93
"succs (Load idx) pc         = {pc+1}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    94
"succs (Store idx) pc        = {pc+1}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    95
"succs (Bipush i) pc         = {pc+1}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    96
"succs (Aconst_null) pc      = {pc+1}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    97
"succs (Getfield F C) pc     = {pc+1}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    98
"succs (Putfield F C) pc     = {pc+1}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    99
"succs (New C) pc            = {pc+1}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   100
"succs (Checkcast C) pc      = {pc+1}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   101
"succs Pop pc                = {pc+1}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   102
"succs Dup pc                = {pc+1}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   103
"succs Dup_x1 pc             = {pc+1}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   104
"succs Dup_x2 pc             = {pc+1}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   105
"succs Swap pc               = {pc+1}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   106
"succs IAdd pc               = {pc+1}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   107
"succs (Ifcmpeq b) pc        = {pc+1, nat (int pc + b)}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   108
"succs (Goto b) pc           = {nat (int pc + b)}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   109
"succs Return pc             = {}"  
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   110
"succs (Invoke C mn fpTs) pc = {pc+1}"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   111
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   112
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   113
lemma 1: "2 < length a \\<Longrightarrow> (\\<exists>l l' l'' ls. a = l#l'#l''#ls)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   114
proof (cases a)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   115
  fix x xs assume "a = x#xs" "2 < length a"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   116
  thus ?thesis by - (cases xs, simp, cases "tl xs", auto)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   117
qed auto
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   118
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   119
lemma 2: "\\<not>(2 < length a) \\<Longrightarrow> a = [] \\<or> (\\<exists> l. a = [l]) \\<or> (\\<exists> l l'. a = [l,l'])"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   120
proof -;
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   121
  assume "\\<not>(2 < length a)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   122
  hence "length a < (Suc 2)" by simp
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   123
  hence * : "length a = 0 \\<or> length a = 1 \\<or> length a = 2" by (auto simp add: less_Suc_eq)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   124
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   125
  { 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   126
    fix x 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   127
    assume "length x = 1"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   128
    hence "\\<exists> l. x = [l]"  by - (cases x, auto)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   129
  } note 0 = this
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   130
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   131
  have "length a = 2 \\<Longrightarrow> \\<exists>l l'. a = [l,l']" by (cases a, auto dest: 0)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   132
  with * show ?thesis by (auto dest: 0)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   133
qed
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   134
9559
kleing
parents: 9549
diff changeset
   135
text {* 
kleing
parents: 9549
diff changeset
   136
\mdeskip
kleing
parents: 9549
diff changeset
   137
simp rules for \isa{app} without patterns, better suited for proofs
kleing
parents: 9549
diff changeset
   138
\medskip
kleing
parents: 9549
diff changeset
   139
*}
kleing
parents: 9549
diff changeset
   140
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   141
lemma appStore[simp]:
9559
kleing
parents: 9549
diff changeset
   142
"app (Store idx, G, rT, s) = (\\<exists> ts ST LT. s = (ts#ST,LT) \\<and> idx < length LT)"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   143
by (cases s, cases "2 < length (fst s)", auto dest: 1 2)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   144
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   145
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   146
lemma appGetField[simp]:
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   147
"app (Getfield F C, G, rT, s) = (\\<exists> oT ST LT. s = (oT#ST, LT) \\<and> is_class G C \\<and> 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   148
                                 fst (the (field (G,C) F)) = C \\<and>
9559
kleing
parents: 9549
diff changeset
   149
                                 field (G,C) F \\<noteq> None \\<and> G \\<turnstile> oT \\<preceq> (Class C))"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   150
by (cases s, cases "2 < length (fst s)", auto dest: 1 2)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   151
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   152
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   153
lemma appPutField[simp]:
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   154
"app (Putfield F C, G, rT, s) = (\\<exists> vT oT ST LT. s = (vT#oT#ST, LT) \\<and> is_class G C \\<and> 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   155
                                 field (G,C) F \\<noteq> None \\<and> fst (the (field (G,C) F)) = C \\<and>
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   156
                                 G \\<turnstile> oT \\<preceq> (Class C) \\<and>
9559
kleing
parents: 9549
diff changeset
   157
                                 G \\<turnstile> vT \\<preceq> (snd (the (field (G,C) F))))" 
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   158
by (cases s, cases "2 < length (fst s)", auto dest: 1 2)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   159
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   160
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   161
lemma appCheckcast[simp]:
9559
kleing
parents: 9549
diff changeset
   162
"app (Checkcast C, G, rT, s) = (\\<exists>rT ST LT. s = (RefT rT#ST,LT) \\<and> is_class G C)"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   163
by (cases s, cases "fst s", simp, cases "hd (fst s)", auto)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   164
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   165
lemma appPop[simp]:
9559
kleing
parents: 9549
diff changeset
   166
"app (Pop, G, rT, s) = (\\<exists>ts ST LT. s = (ts#ST,LT))" 
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   167
by (cases s, cases "2 < length (fst s)", auto dest: 1 2)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   168
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   169
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   170
lemma appDup[simp]:
9559
kleing
parents: 9549
diff changeset
   171
"app (Dup, G, rT, s) = (\\<exists>ts ST LT. s = (ts#ST,LT))" 
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   172
by (cases s, cases "2 < length (fst s)", auto dest: 1 2)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   173
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   174
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   175
lemma appDup_x1[simp]:
9559
kleing
parents: 9549
diff changeset
   176
"app (Dup_x1, G, rT, s) = (\\<exists>ts1 ts2 ST LT. s = (ts1#ts2#ST,LT))" 
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   177
by (cases s, cases "2 < length (fst s)", auto dest: 1 2)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   178
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   179
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   180
lemma appDup_x2[simp]:
9559
kleing
parents: 9549
diff changeset
   181
"app (Dup_x2, G, rT, s) = (\\<exists>ts1 ts2 ts3 ST LT. s = (ts1#ts2#ts3#ST,LT))"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   182
by (cases s, cases "2 < length (fst s)", auto dest: 1 2)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   183
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   184
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   185
lemma appSwap[simp]:
9559
kleing
parents: 9549
diff changeset
   186
"app (Swap, G, rT, s) = (\\<exists>ts1 ts2 ST LT. s = (ts1#ts2#ST,LT))" 
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   187
by (cases s, cases "2 < length (fst s)", auto dest: 1 2)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   188
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   189
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   190
lemma appIAdd[simp]:
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   191
"app (IAdd, G, rT, s) = (\\<exists> ST LT. s = (PrimT Integer#PrimT Integer#ST,LT))"  (is "?app s = ?P s")
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   192
proof (cases s)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   193
  case Pair
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   194
  have "?app (a,b) = ?P (a,b)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   195
  proof (cases "a")
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   196
    fix t ts assume a: "a = t#ts"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   197
    show ?thesis
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   198
    proof (cases t)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   199
      fix p assume p: "t = PrimT p"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   200
      show ?thesis
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   201
      proof (cases p)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   202
        assume ip: "p = Integer"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   203
        show ?thesis
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   204
        proof (cases ts)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   205
          fix t' ts' assume t': "ts = t' # ts'"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   206
          show ?thesis
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   207
          proof (cases t')
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   208
            fix p' assume "t' = PrimT p'"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   209
            with t' ip p a
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   210
            show ?thesis by - (cases p', auto)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   211
          qed (auto simp add: a p ip t')
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   212
        qed (auto simp add: a p ip)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   213
      qed (auto simp add: a p)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   214
    qed (auto simp add: a)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   215
  qed auto
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   216
  with Pair show ?thesis by simp
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   217
qed
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   218
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   219
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   220
lemma appIfcmpeq[simp]:
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   221
"app (Ifcmpeq b, G, rT, s) = (\\<exists>ts1 ts2 ST LT. s = (ts1#ts2#ST,LT) \\<and> 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   222
                              ((\\<exists> p. ts1 = PrimT p \\<and> ts1 = PrimT p) \\<or>  
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   223
                               (\\<exists>r r'. ts1 = RefT r \\<and> ts2 = RefT r')))" 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   224
by (cases s, cases "2 < length (fst s)", auto dest: 1 2)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   225
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   226
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   227
lemma appReturn[simp]:
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   228
"app (Return, G, rT, s) = (\\<exists>T ST LT. s = (T#ST,LT) \\<and> (G \\<turnstile> T \\<preceq> rT))" 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   229
by (cases s, cases "2 < length (fst s)", auto dest: 1 2)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   230
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   231
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   232
lemma appInvoke[simp]:
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   233
"app (Invoke C mn fpTs, G, rT, s) = (\\<exists>apTs X ST LT.
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   234
                                       s = ((rev apTs) @ (X # ST), LT) \\<and> 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   235
                                       length apTs = length fpTs \\<and> 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   236
                                       G \\<turnstile> X \\<preceq> Class C \\<and>  
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   237
                                       (\\<forall>(aT,fT)\\<in>set(zip apTs fpTs). G \\<turnstile> aT \\<preceq> fT) \\<and> 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   238
                                       method (G,C) (mn,fpTs) \\<noteq> None)" (is "?app s = ?P s")
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   239
proof (cases s)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   240
  case Pair
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   241
  have "?app (a,b) \\<Longrightarrow> ?P (a,b)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   242
  proof -
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   243
    assume app: "?app (a,b)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   244
    hence "a = (rev (rev (take (length fpTs) a))) @ (drop (length fpTs) a) \\<and> length fpTs < length a" 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   245
      (is "?a \\<and> ?l") by auto    
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   246
    hence "?a \\<and> 0 < length (drop (length fpTs) a)" (is "?a \\<and> ?l") by auto
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   247
    hence "?a \\<and> ?l \\<and> length (rev (take (length fpTs) a)) = length fpTs" by (auto simp add: min_def)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   248
    hence "\\<exists>apTs ST. a = rev apTs @ ST \\<and> length apTs = length fpTs \\<and> 0 < length ST" by blast
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   249
    hence "\\<exists>apTs ST. a = rev apTs @ ST \\<and> length apTs = length fpTs \\<and> ST \\<noteq> []" by blast        
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   250
    hence "\\<exists>apTs ST. a = rev apTs @ ST \\<and> length apTs = length fpTs \\<and> (\\<exists>X ST'. ST = X#ST')" by (simp add: neq_Nil_conv)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   251
    hence "\\<exists>apTs X ST. a = rev apTs @ X # ST \\<and> length apTs = length fpTs" by blast
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   252
    with app
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   253
    show ?thesis by auto blast
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   254
  qed
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   255
  with Pair have "?app s \\<Longrightarrow> ?P s" by simp
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   256
  thus ?thesis by auto
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   257
qed 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   258
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   259
lemmas [simp del] = app_invoke
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   260
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   261
end