src/HOL/MicroJava/BV/Step.thy
author nipkow
Mon, 06 Aug 2001 15:54:29 +0200
changeset 11466 d64ccdeaf9ae
parent 10897 697fab84709e
child 11701 3d51fbf81c17
permissions -rw-r--r--
1 -> 1'
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
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10623
diff changeset
    10
theory Step = JVMType + JVMExec:
9549
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
9594
42d11e0a7a8b Convert.thy now in Isar, tuned
kleing
parents: 9585
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 
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9757
diff changeset
    15
step' :: "instr \<times> jvm_prog \<times> state_type => state_type"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    16
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    17
recdef step' "{}"
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
    18
"step' (Load idx,  G, (ST, LT))          = (ok_val (LT ! idx) # ST, LT)"
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
    19
"step' (Store idx, G, (ts#ST, LT))       = (ST, LT[idx:= OK ts])"
10897
697fab84709e removed instructions Aconst_null+Bipush, introduced LitPush
kleing
parents: 10812
diff changeset
    20
"step' (LitPush v, G, (ST, LT))           = (the (typeof (\<lambda>v. None) v) # ST, LT)"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    21
"step' (Getfield F C, G, (oT#ST, LT))    = (snd (the (field (G,C) F)) # ST, LT)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    22
"step' (Putfield F C, G, (vT#oT#ST, LT)) = (ST,LT)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    23
"step' (New C, G, (ST,LT))               = (Class C # ST, LT)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    24
"step' (Checkcast C, G, (RefT rt#ST,LT)) = (Class C # ST,LT)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    25
"step' (Pop, G, (ts#ST,LT))              = (ST,LT)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    26
"step' (Dup, G, (ts#ST,LT))              = (ts#ts#ST,LT)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    27
"step' (Dup_x1, G, (ts1#ts2#ST,LT))      = (ts1#ts2#ts1#ST,LT)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    28
"step' (Dup_x2, G, (ts1#ts2#ts3#ST,LT))  = (ts1#ts2#ts3#ts1#ST,LT)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    29
"step' (Swap, G, (ts1#ts2#ST,LT))        = (ts2#ts1#ST,LT)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    30
"step' (IAdd, G, (PrimT Integer#PrimT Integer#ST,LT)) 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    31
                                         = (PrimT Integer#ST,LT)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    32
"step' (Ifcmpeq b, G, (ts1#ts2#ST,LT))   = (ST,LT)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    33
"step' (Goto b, G, s)                    = s"
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
    34
  (* Return has no successor instruction in the same method *)
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
    35
"step' (Return, G, s)                    = s" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    36
"step' (Invoke C mn fpTs, G, (ST,LT))    = (let ST' = drop (length fpTs) ST 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    37
  in  (fst (snd (the (method (G,C) (mn,fpTs))))#(tl ST'),LT))" 
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    38
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    39
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    40
constdefs
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9757
diff changeset
    41
  step :: "instr => jvm_prog => state_type option => state_type option"
7164dc0d24d8 unsymbolized
kleing
parents: 9757
diff changeset
    42
  "step i G == option_map (\<lambda>s. step' (i,G,s))"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    43
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    44
9594
42d11e0a7a8b Convert.thy now in Isar, tuned
kleing
parents: 9585
diff changeset
    45
text "Conditions under which step is applicable:"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    46
consts
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    47
app' :: "instr \<times> jvm_prog \<times> nat \<times> ty \<times> state_type => bool"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    48
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    49
recdef app' "{}"
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    50
"app' (Load idx, G, maxs, rT, s)                  = (idx < length (snd s) \<and> 
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    51
                                                    (snd s) ! idx \<noteq> Err \<and> 
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    52
                                                    maxs < length (fst s))"
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    53
"app' (Store idx, G, maxs, rT, (ts#ST, LT))       = (idx < length LT)"
10897
697fab84709e removed instructions Aconst_null+Bipush, introduced LitPush
kleing
parents: 10812
diff changeset
    54
"app' (LitPush v, G, maxs, rT, s)                  = (maxs < length (fst s) \<and> typeof (\<lambda>t. None) v \<noteq> None)"
697fab84709e removed instructions Aconst_null+Bipush, introduced LitPush
kleing
parents: 10812
diff changeset
    55
"app' (Getfield F C, G, maxs, rT, (oT#ST, LT))    = (is_class G C \<and>
9580
kleing
parents: 9559
diff changeset
    56
                                              field (G,C) F \<noteq> None \<and>
kleing
parents: 9559
diff changeset
    57
                                              fst (the (field (G,C) F)) = C \<and>
kleing
parents: 9559
diff changeset
    58
                                              G \<turnstile> oT \<preceq> (Class C))"
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    59
"app' (Putfield F C, G, maxs, rT, (vT#oT#ST, LT)) = (is_class G C \<and> 
9580
kleing
parents: 9559
diff changeset
    60
                                              field (G,C) F \<noteq> None \<and>
kleing
parents: 9559
diff changeset
    61
                                              fst (the (field (G,C) F)) = C \<and>
kleing
parents: 9559
diff changeset
    62
                                              G \<turnstile> oT \<preceq> (Class C) \<and>
kleing
parents: 9559
diff changeset
    63
                                              G \<turnstile> vT \<preceq> (snd (the (field (G,C) F))))" 
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    64
"app' (New C, G, maxs, rT, s)                     = (is_class G C \<and> maxs < length (fst s))"
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    65
"app' (Checkcast C, G, maxs, rT, (RefT rt#ST,LT)) = (is_class G C)"
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    66
"app' (Pop, G, maxs, rT, (ts#ST,LT))              = True"
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    67
"app' (Dup, G, maxs, rT, (ts#ST,LT))              = (maxs < Suc (length ST))"
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    68
"app' (Dup_x1, G, maxs, rT, (ts1#ts2#ST,LT))      = (maxs < Suc (Suc (length ST)))"
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    69
"app' (Dup_x2, G, maxs, rT, (ts1#ts2#ts3#ST,LT))  = (maxs < Suc (Suc (Suc (length ST))))"
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    70
"app' (Swap, G, maxs, rT, (ts1#ts2#ST,LT))        = True"
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    71
"app' (IAdd, G, maxs, rT, (PrimT Integer#PrimT Integer#ST,LT)) 
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    72
                                                 = True"
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    73
"app' (Ifcmpeq b, G, maxs, rT, (ts#ts'#ST,LT))    = ((\<exists>p. ts = PrimT p \<and> ts' = PrimT p) \<or> 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    74
                                              (\<exists>r r'. ts = RefT r \<and> ts' = RefT r'))"
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    75
"app' (Goto b, G, maxs, rT, s)                    = True"
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    76
"app' (Return, G, maxs, rT, (T#ST,LT))            = (G \<turnstile> T \<preceq> rT)"
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    77
"app' (Invoke C mn fpTs, G, maxs, rT, s)          = 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    78
   (length fpTs < length (fst s) \<and> 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    79
   (let apTs = rev (take (length fpTs) (fst s));
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    80
        X    = hd (drop (length fpTs) (fst s)) 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    81
    in  
10623
f16baa9505cd invoked class must be defined in Invoke C ...
kleing
parents: 10592
diff changeset
    82
        G \<turnstile> X \<preceq> Class C \<and> is_class G C \<and> method (G,C) (mn,fpTs) \<noteq> None \<and> 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    83
        (\<forall>(aT,fT)\<in>set(zip apTs fpTs). G \<turnstile> aT \<preceq> fT)))"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    84
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    85
"app' (i,G,maxs,rT,s)                             = False"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    86
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    87
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    88
constdefs
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    89
  app :: "instr => jvm_prog => nat => ty => state_type option => bool"
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
    90
  "app i G maxs rT s == case s of None => True | Some t => app' (i,G,maxs,rT,t)"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    91
9594
42d11e0a7a8b Convert.thy now in Isar, tuned
kleing
parents: 9585
diff changeset
    92
text {* program counter of successor instructions: *}
9559
kleing
parents: 9549
diff changeset
    93
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    94
consts
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9757
diff changeset
    95
succs :: "instr => p_count => p_count list"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    96
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
    97
primrec 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    98
"succs (Load idx) pc         = [pc+1]"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    99
"succs (Store idx) pc        = [pc+1]"
10897
697fab84709e removed instructions Aconst_null+Bipush, introduced LitPush
kleing
parents: 10812
diff changeset
   100
"succs (LitPush v) pc        = [pc+1]"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   101
"succs (Getfield F C) pc     = [pc+1]"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   102
"succs (Putfield F C) pc     = [pc+1]"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   103
"succs (New C) pc            = [pc+1]"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   104
"succs (Checkcast C) pc      = [pc+1]"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   105
"succs Pop pc                = [pc+1]"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   106
"succs Dup pc                = [pc+1]"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   107
"succs Dup_x1 pc             = [pc+1]"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   108
"succs Dup_x2 pc             = [pc+1]"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   109
"succs Swap pc               = [pc+1]"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   110
"succs IAdd pc               = [pc+1]"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   111
"succs (Ifcmpeq b) pc        = [pc+1, nat (int pc + b)]"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   112
"succs (Goto b) pc           = [nat (int pc + b)]"
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
   113
"succs Return pc             = [pc]"  
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   114
"succs (Invoke C mn fpTs) pc = [pc+1]"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   115
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   116
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9757
diff changeset
   117
lemma 1: "2 < length a ==> (\<exists>l l' l'' ls. a = l#l'#l''#ls)"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   118
proof (cases a)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   119
  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
   120
  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
   121
qed auto
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   122
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9757
diff changeset
   123
lemma 2: "\<not>(2 < length a) ==> a = [] \<or> (\<exists> l. a = [l]) \<or> (\<exists> l l'. a = [l,l'])"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   124
proof -;
9580
kleing
parents: 9559
diff changeset
   125
  assume "\<not>(2 < length a)"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   126
  hence "length a < (Suc 2)" by simp
11466
d64ccdeaf9ae 1 -> 1'
nipkow
parents: 10897
diff changeset
   127
  hence * : "length a = 0 \<or> length a = 1' \<or> length a = 2" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   128
    by (auto simp add: less_Suc_eq)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   129
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
    fix x 
11466
d64ccdeaf9ae 1 -> 1'
nipkow
parents: 10897
diff changeset
   132
    assume "length x = 1'"
9580
kleing
parents: 9559
diff changeset
   133
    hence "\<exists> l. x = [l]"  by - (cases x, auto)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   134
  } note 0 = this
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   135
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9757
diff changeset
   136
  have "length a = 2 ==> \<exists>l l'. a = [l,l']" by (cases a, auto dest: 0)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   137
  with * show ?thesis by (auto dest: 0)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   138
qed
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   139
9559
kleing
parents: 9549
diff changeset
   140
text {* 
9585
f0e811a54254 fixed document preparation;
wenzelm
parents: 9580
diff changeset
   141
\medskip
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   142
simp rules for @{term app}
9559
kleing
parents: 9549
diff changeset
   143
*}
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   144
lemma appNone[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   145
"app i G maxs rT None = True"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   146
  by (simp add: app_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   147
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   148
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   149
lemma appLoad[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   150
"(app (Load idx) G maxs rT (Some s)) = (idx < length (snd s) \<and> (snd s) ! idx \<noteq> Err \<and> maxs < length (fst s))"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   151
  by (simp add: app_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   152
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   153
lemma appStore[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   154
"(app (Store idx) G maxs rT (Some s)) = (\<exists> ts ST LT. s = (ts#ST,LT) \<and> idx < length LT)"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   155
  by (cases s, cases "2 < length (fst s)", auto dest: 1 2 simp add: app_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   156
10897
697fab84709e removed instructions Aconst_null+Bipush, introduced LitPush
kleing
parents: 10812
diff changeset
   157
lemma appLitPush[simp]:
697fab84709e removed instructions Aconst_null+Bipush, introduced LitPush
kleing
parents: 10812
diff changeset
   158
"(app (LitPush v) G maxs rT (Some s)) = (maxs < length (fst s) \<and> typeof (\<lambda>v. None) v \<noteq> None)"
697fab84709e removed instructions Aconst_null+Bipush, introduced LitPush
kleing
parents: 10812
diff changeset
   159
  by (cases s, simp add: app_def)
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   160
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   161
lemma appGetField[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   162
"(app (Getfield F C) G maxs rT (Some s)) = 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   163
 (\<exists> oT vT ST LT. s = (oT#ST, LT) \<and> is_class G C \<and>  
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   164
  field (G,C) F = Some (C,vT) \<and> G \<turnstile> oT \<preceq> (Class C))"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   165
  by (cases s, cases "2 < length (fst s)", auto dest!: 1 2 simp add: app_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   166
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   167
lemma appPutField[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   168
"(app (Putfield F C) G maxs rT (Some s)) = 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   169
 (\<exists> vT vT' oT ST LT. s = (vT#oT#ST, LT) \<and> is_class G C \<and> 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   170
  field (G,C) F = Some (C, vT') \<and> G \<turnstile> oT \<preceq> (Class C) \<and> G \<turnstile> vT \<preceq> vT')"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   171
  by (cases s, cases "2 < length (fst s)", auto dest!: 1 2 simp add: app_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   172
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   173
lemma appNew[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   174
"(app (New C) G maxs rT (Some s)) = (is_class G C \<and> maxs < length (fst s))"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   175
  by (simp add: app_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   176
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   177
lemma appCheckcast[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   178
"(app (Checkcast C) G maxs rT (Some s)) = (\<exists>rT ST LT. s = (RefT rT#ST,LT) \<and> is_class G C)"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   179
  by (cases s, cases "fst s", simp add: app_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   180
     (cases "hd (fst s)", auto simp add: app_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   181
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   182
lemma appPop[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   183
"(app Pop G maxs rT (Some s)) = (\<exists>ts ST LT. s = (ts#ST,LT))" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   184
  by (cases s, cases "2 < length (fst s)", auto dest: 1 2 simp add: app_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   185
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   186
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   187
lemma appDup[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   188
"(app Dup G maxs rT (Some s)) = (\<exists>ts ST LT. s = (ts#ST,LT) \<and> maxs < Suc (length ST))" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   189
  by (cases s, cases "2 < length (fst s)", auto dest: 1 2 simp add: app_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   190
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   191
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   192
lemma appDup_x1[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   193
"(app Dup_x1 G maxs rT (Some s)) = (\<exists>ts1 ts2 ST LT. s = (ts1#ts2#ST,LT) \<and> maxs < Suc (Suc (length ST)))" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   194
  by (cases s, cases "2 < length (fst s)", auto dest: 1 2 simp add: app_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   195
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   196
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   197
lemma appDup_x2[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   198
"(app Dup_x2 G maxs rT (Some s)) = (\<exists>ts1 ts2 ts3 ST LT. s = (ts1#ts2#ts3#ST,LT) \<and> maxs < Suc (Suc (Suc (length ST))))"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   199
  by (cases s, cases "2 < length (fst s)", auto dest: 1 2 simp add: app_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   200
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   201
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   202
lemma appSwap[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   203
"app Swap G maxs rT (Some s) = (\<exists>ts1 ts2 ST LT. s = (ts1#ts2#ST,LT))" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   204
  by (cases s, cases "2 < length (fst s)", auto dest: 1 2 simp add: app_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   205
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   206
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   207
lemma appIAdd[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   208
"app IAdd G maxs rT (Some s) = (\<exists> ST LT. s = (PrimT Integer#PrimT Integer#ST,LT))"  
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   209
  (is "?app s = ?P s")
9664
4cae97480a6d open cases;
wenzelm
parents: 9594
diff changeset
   210
proof (cases (open) s)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   211
  case Pair
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   212
  have "?app (a,b) = ?P (a,b)"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   213
  proof (cases "a")
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   214
    fix t ts assume a: "a = t#ts"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   215
    show ?thesis
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   216
    proof (cases t)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   217
      fix p assume p: "t = PrimT p"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   218
      show ?thesis
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   219
      proof (cases p)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   220
        assume ip: "p = Integer"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   221
        show ?thesis
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   222
        proof (cases ts)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   223
          fix t' ts' assume t': "ts = t' # ts'"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   224
          show ?thesis
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   225
          proof (cases t')
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   226
            fix p' assume "t' = PrimT p'"
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   227
            with t' ip p a
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   228
            show ?thesis by - (cases p', auto simp add: app_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   229
          qed (auto simp add: a p ip t' app_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   230
        qed (auto simp add: a p ip app_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   231
      qed (auto simp add: a p app_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   232
    qed (auto simp add: a app_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   233
  qed (auto simp add: app_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   234
  with Pair show ?thesis by simp
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   235
qed
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   236
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   237
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   238
lemma appIfcmpeq[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   239
"app (Ifcmpeq b) G maxs rT (Some s) = (\<exists>ts1 ts2 ST LT. s = (ts1#ts2#ST,LT) \<and> 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   240
 ((\<exists> p. ts1 = PrimT p \<and> ts2 = PrimT p) \<or> (\<exists>r r'. ts1 = RefT r \<and> ts2 = RefT r')))" 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   241
  by (cases s, cases "2 < length (fst s)", auto dest: 1 2 simp add: app_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   242
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   243
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   244
lemma appReturn[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   245
"app Return G maxs rT (Some s) = (\<exists>T ST LT. s = (T#ST,LT) \<and> (G \<turnstile> T \<preceq> rT))" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   246
  by (cases s, cases "2 < length (fst s)", auto dest: 1 2 simp add: app_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   247
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   248
lemma appGoto[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   249
"app (Goto branch) G maxs rT (Some s) = True"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   250
  by (simp add: app_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   251
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   252
lemma appInvoke[simp]:
10592
fc0b575a2cf7 BCV Integration
kleing
parents: 10496
diff changeset
   253
"app (Invoke C mn fpTs) G maxs rT (Some s) = (\<exists>apTs X ST LT mD' rT' b'.
10623
f16baa9505cd invoked class must be defined in Invoke C ...
kleing
parents: 10592
diff changeset
   254
  s = ((rev apTs) @ (X # ST), LT) \<and> length apTs = length fpTs \<and> is_class G C \<and>
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   255
  G \<turnstile> X \<preceq> Class C \<and> (\<forall>(aT,fT)\<in>set(zip apTs fpTs). G \<turnstile> aT \<preceq> fT) \<and> 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   256
  method (G,C) (mn,fpTs) = Some (mD', rT', b'))" (is "?app s = ?P s")
9664
4cae97480a6d open cases;
wenzelm
parents: 9594
diff changeset
   257
proof (cases (open) s)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   258
  case Pair
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9757
diff changeset
   259
  have "?app (a,b) ==> ?P (a,b)"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   260
  proof -
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   261
    assume app: "?app (a,b)"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   262
    hence "a = (rev (rev (take (length fpTs) a))) @ (drop (length fpTs) a) \<and> 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   263
           length fpTs < length a" (is "?a \<and> ?l") 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   264
      by (auto simp add: app_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   265
    hence "?a \<and> 0 < length (drop (length fpTs) a)" (is "?a \<and> ?l") 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   266
      by auto
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   267
    hence "?a \<and> ?l \<and> length (rev (take (length fpTs) a)) = length fpTs" 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   268
      by (auto simp add: min_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   269
    hence "\<exists>apTs ST. a = rev apTs @ ST \<and> length apTs = length fpTs \<and> 0 < length ST" 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   270
      by blast
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   271
    hence "\<exists>apTs ST. a = rev apTs @ ST \<and> length apTs = length fpTs \<and> ST \<noteq> []" 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   272
      by blast
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   273
    hence "\<exists>apTs ST. a = rev apTs @ ST \<and> length apTs = length fpTs \<and> 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   274
           (\<exists>X ST'. ST = X#ST')" 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   275
      by (simp add: neq_Nil_conv)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   276
    hence "\<exists>apTs X ST. a = rev apTs @ X # ST \<and> length apTs = length fpTs" 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   277
      by blast
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   278
    with app
10897
697fab84709e removed instructions Aconst_null+Bipush, introduced LitPush
kleing
parents: 10812
diff changeset
   279
    show ?thesis 
697fab84709e removed instructions Aconst_null+Bipush, introduced LitPush
kleing
parents: 10812
diff changeset
   280
      by (unfold app_def, clarsimp) blast
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   281
  qed
10897
697fab84709e removed instructions Aconst_null+Bipush, introduced LitPush
kleing
parents: 10812
diff changeset
   282
  with Pair 
697fab84709e removed instructions Aconst_null+Bipush, introduced LitPush
kleing
parents: 10812
diff changeset
   283
  have "?app s ==> ?P s" by simp
697fab84709e removed instructions Aconst_null+Bipush, introduced LitPush
kleing
parents: 10812
diff changeset
   284
  moreover
697fab84709e removed instructions Aconst_null+Bipush, introduced LitPush
kleing
parents: 10812
diff changeset
   285
  have "?P s \<Longrightarrow> ?app s" by (unfold app_def) clarsimp
697fab84709e removed instructions Aconst_null+Bipush, introduced LitPush
kleing
parents: 10812
diff changeset
   286
  ultimately
697fab84709e removed instructions Aconst_null+Bipush, introduced LitPush
kleing
parents: 10812
diff changeset
   287
  show ?thesis by blast
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   288
qed 
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   289
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   290
lemma step_Some:
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   291
  "step i G (Some s) \<noteq> None"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   292
  by (simp add: step_def)
9580
kleing
parents: 9559
diff changeset
   293
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   294
lemma step_None [simp]:
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   295
  "step i G None = None"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   296
  by (simp add: step_def)
9580
kleing
parents: 9559
diff changeset
   297
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents:
diff changeset
   298
end