| author | haftmann | 
| Sun, 26 Jul 2009 08:03:40 +0200 | |
| changeset 32208 | e6a42620e6c1 | 
| parent 30952 | 7ab2716dd93b | 
| child 32960 | 69916a850301 | 
| permissions | -rw-r--r-- | 
| 12857 | 1 | (* Title: HOL/Bali/Trans.thy | 
| 12854 | 2 | ID: $Id$ | 
| 13688 
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
 schirmer parents: 
13384diff
changeset | 3 | Author: David von Oheimb and Norbert Schirmer | 
| 12854 | 4 | |
| 5 | Operational transition (small-step) semantics of the | |
| 6 | execution of Java expressions and statements | |
| 7 | ||
| 8 | PRELIMINARY!!!!!!!! | |
| 9 | *) | |
| 10 | ||
| 16417 | 11 | theory Trans imports Evaln begin | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 12 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 13 | constdefs groundVar:: "var \<Rightarrow> bool" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 14 | "groundVar v \<equiv> (case v of | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 15 | LVar ln \<Rightarrow> True | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 16 |                  | {accC,statDeclC,stat}e..fn \<Rightarrow> \<exists> a. e=Lit a
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 17 | | e1.[e2] \<Rightarrow> \<exists> a i. e1= Lit a \<and> e2 = Lit i | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 18 | | InsInitV c v \<Rightarrow> False)" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 19 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 20 | lemma groundVar_cases [consumes 1, case_names LVar FVar AVar]: | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 21 | assumes ground: "groundVar v" and | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 22 | LVar: "\<And> ln. \<lbrakk>v=LVar ln\<rbrakk> \<Longrightarrow> P" and | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 23 | FVar: "\<And> accC statDeclC stat a fn. | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 24 |                     \<lbrakk>v={accC,statDeclC,stat}(Lit a)..fn\<rbrakk> \<Longrightarrow> P" and
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 25 | AVar: "\<And> a i. \<lbrakk>v=(Lit a).[Lit i]\<rbrakk> \<Longrightarrow> P" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 26 | shows "P" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 27 | proof - | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 28 | from ground LVar FVar AVar | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 29 | show ?thesis | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 30 | apply (cases v) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 31 | apply (simp add: groundVar_def) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 32 | apply (simp add: groundVar_def,blast) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 33 | apply (simp add: groundVar_def,blast) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 34 | apply (simp add: groundVar_def) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 35 | done | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 36 | qed | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 37 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 38 | constdefs groundExprs:: "expr list \<Rightarrow> bool" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 39 | "groundExprs es \<equiv> list_all (\<lambda> e. \<exists> v. e=Lit v) es" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 40 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 41 | consts the_val:: "expr \<Rightarrow> val" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 42 | primrec | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 43 | "the_val (Lit v) = v" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 44 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 45 | consts the_var:: "prog \<Rightarrow> state \<Rightarrow> var \<Rightarrow> (vvar \<times> state)" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 46 | primrec | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 47 | "the_var G s (LVar ln) =(lvar ln (store s),s)" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 48 | the_var_FVar_def: | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 49 | "the_var G s ({accC,statDeclC,stat}a..fn) =fvar statDeclC stat fn (the_val a) s"
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 50 | the_var_AVar_def: | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 51 | "the_var G s(a.[i]) =avar G (the_val i) (the_val a) s" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 52 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 53 | lemma the_var_FVar_simp[simp]: | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 54 | "the_var G s ({accC,statDeclC,stat}(Lit a)..fn) = fvar statDeclC stat fn a s"
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 55 | by (simp) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 56 | declare the_var_FVar_def [simp del] | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 57 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 58 | lemma the_var_AVar_simp: | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 59 | "the_var G s ((Lit a).[Lit i]) = avar G i a s" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 60 | by (simp) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 61 | declare the_var_AVar_def [simp del] | 
| 12854 | 62 | |
| 21765 | 63 | syntax (xsymbols) | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 64 | Ref :: "loc \<Rightarrow> expr" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 65 | SKIP :: "expr" | 
| 12854 | 66 | |
| 67 | translations | |
| 68 | "Ref a" == "Lit (Addr a)" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 69 | "SKIP" == "Lit Unit" | 
| 12854 | 70 | |
| 23747 | 71 | inductive | 
| 21765 | 72 |   step :: "[prog,term \<times> state,term \<times> state] \<Rightarrow> bool" ("_\<turnstile>_ \<mapsto>1 _"[61,82,82] 81)
 | 
| 73 | for G :: prog | |
| 74 | where | |
| 12854 | 75 | |
| 76 | (* evaluation of expression *) | |
| 77 | (* cf. 15.5 *) | |
| 21765 | 78 | Abrupt: "\<lbrakk>\<forall>v. t \<noteq> \<langle>Lit v\<rangle>; | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 79 | \<forall> t. t \<noteq> \<langle>l\<bullet> Skip\<rangle>; | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 80 | \<forall> C vn c. t \<noteq> \<langle>Try Skip Catch(C vn) c\<rangle>; | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 81 | \<forall> x c. t \<noteq> \<langle>Skip Finally c\<rangle> \<and> xc \<noteq> Xcpt x; | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 82 | \<forall> a c. t \<noteq> \<langle>FinA a c\<rangle>\<rbrakk> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 83 | \<Longrightarrow> | 
| 28524 | 84 | G\<turnstile>(t,Some xc,s) \<mapsto>1 (\<langle>Lit undefined\<rangle>,Some xc,s)" | 
| 12854 | 85 | |
| 21765 | 86 | | InsInitE: "\<lbrakk>G\<turnstile>(\<langle>c\<rangle>,Norm s) \<mapsto>1 (\<langle>c'\<rangle>, s')\<rbrakk> | 
| 87 | \<Longrightarrow> | |
| 88 | G\<turnstile>(\<langle>InsInitE c e\<rangle>,Norm s) \<mapsto>1 (\<langle>InsInitE c' e\<rangle>, s')" | |
| 12854 | 89 | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 90 | (* SeqE: "G\<turnstile>(\<langle>Seq Skip e\<rangle>,Norm s) \<mapsto>1 (\<langle>e\<rangle>, Norm s)" *) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 91 | (* Specialised rules to evaluate: | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 92 | InsInitE Skip (NewC C), InisInitE Skip (NewA T[e]) *) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 93 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 94 | (* cf. 15.8.1 *) | 
| 21765 | 95 | | NewC: "G\<turnstile>(\<langle>NewC C\<rangle>,Norm s) \<mapsto>1 (\<langle>InsInitE (Init C) (NewC C)\<rangle>, Norm s)" | 
| 96 | | NewCInited: "\<lbrakk>G\<turnstile> Norm s \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s'\<rbrakk> | |
| 97 | \<Longrightarrow> | |
| 98 | G\<turnstile>(\<langle>InsInitE Skip (NewC C)\<rangle>,Norm s) \<mapsto>1 (\<langle>Ref a\<rangle>, s')" | |
| 12854 | 99 | |
| 100 | ||
| 101 | ||
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 102 | (* Alternative when rule SeqE is present | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 103 | NewCInited: "\<lbrakk>inited C (globs s); | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 104 | G\<turnstile> Norm s \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s'\<rbrakk> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 105 | \<Longrightarrow> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 106 | G\<turnstile>(\<langle>NewC C\<rangle>,Norm s) \<mapsto>1 (\<langle>Ref a\<rangle>, s')" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 107 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 108 | NewC: | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 109 | "\<lbrakk>\<not> inited C (globs s)\<rbrakk> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 110 | \<Longrightarrow> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 111 | G\<turnstile>(\<langle>NewC C\<rangle>,Norm s) \<mapsto>1 (\<langle>Seq (Init C) (NewC C)\<rangle>, Norm s)" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 112 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 113 | *) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 114 | (* cf. 15.9.1 *) | 
| 21765 | 115 | | NewA: | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 116 | "G\<turnstile>(\<langle>New T[e]\<rangle>,Norm s) \<mapsto>1 (\<langle>InsInitE (init_comp_ty T) (New T[e])\<rangle>,Norm s)" | 
| 21765 | 117 | | InsInitNewAIdx: | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 118 | "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>, s')\<rbrakk> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 119 | \<Longrightarrow> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 120 | G\<turnstile>(\<langle>InsInitE Skip (New T[e])\<rangle>,Norm s) \<mapsto>1 (\<langle>InsInitE Skip (New T[e'])\<rangle>,s')" | 
| 21765 | 121 | | InsInitNewA: | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 122 | "\<lbrakk>G\<turnstile>abupd (check_neg i) (Norm s) \<midarrow>halloc (Arr T (the_Intg i))\<succ>a\<rightarrow> s' \<rbrakk> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 123 | \<Longrightarrow> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 124 | G\<turnstile>(\<langle>InsInitE Skip (New T[Lit i])\<rangle>,Norm s) \<mapsto>1 (\<langle>Ref a\<rangle>,s')" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 125 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 126 | (* cf. 15.15 *) | 
| 21765 | 127 | | CastE: | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 128 | "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 129 | \<Longrightarrow> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 130 | G\<turnstile>(\<langle>Cast T e\<rangle>,None,s) \<mapsto>1 (\<langle>Cast T e'\<rangle>,s')" | 
| 21765 | 131 | | Cast: | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 132 | "\<lbrakk>s' = abupd (raise_if (\<not>G,s\<turnstile>v fits T) ClassCast) (Norm s)\<rbrakk> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 133 | \<Longrightarrow> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 134 | G\<turnstile>(\<langle>Cast T (Lit v)\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit v\<rangle>,s')" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 135 | (* can be written without abupd, since we know Norm s *) | 
| 12854 | 136 | |
| 137 | ||
| 21765 | 138 | | InstE: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'::expr\<rangle>,s')\<rbrakk> | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 139 | \<Longrightarrow> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 140 | G\<turnstile>(\<langle>e InstOf T\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')" | 
| 21765 | 141 | | Inst: "\<lbrakk>b = (v\<noteq>Null \<and> G,s\<turnstile>v fits RefT T)\<rbrakk> | 
| 142 | \<Longrightarrow> | |
| 143 | G\<turnstile>(\<langle>(Lit v) InstOf T\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit (Bool b)\<rangle>,s')" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 144 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 145 | (* cf. 15.7.1 *) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 146 | (*Lit "G\<turnstile>(Lit v,None,s) \<mapsto>1 (Lit v,None,s)"*) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 147 | |
| 21765 | 148 | | UnOpE: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s') \<rbrakk> | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 149 | \<Longrightarrow> | 
| 21765 | 150 | G\<turnstile>(\<langle>UnOp unop e\<rangle>,Norm s) \<mapsto>1 (\<langle>UnOp unop e'\<rangle>,s')" | 
| 151 | | UnOp: "G\<turnstile>(\<langle>UnOp unop (Lit v)\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit (eval_unop unop v)\<rangle>,Norm s)" | |
| 152 | ||
| 153 | | BinOpE1: "\<lbrakk>G\<turnstile>(\<langle>e1\<rangle>,Norm s) \<mapsto>1 (\<langle>e1'\<rangle>,s') \<rbrakk> | |
| 154 | \<Longrightarrow> | |
| 155 | G\<turnstile>(\<langle>BinOp binop e1 e2\<rangle>,Norm s) \<mapsto>1 (\<langle>BinOp binop e1' e2\<rangle>,s')" | |
| 156 | | BinOpE2: "\<lbrakk>need_second_arg binop v1; G\<turnstile>(\<langle>e2\<rangle>,Norm s) \<mapsto>1 (\<langle>e2'\<rangle>,s') \<rbrakk> | |
| 13384 | 157 | \<Longrightarrow> | 
| 158 | G\<turnstile>(\<langle>BinOp binop (Lit v1) e2\<rangle>,Norm s) | |
| 21765 | 159 | \<mapsto>1 (\<langle>BinOp binop (Lit v1) e2'\<rangle>,s')" | 
| 160 | | BinOpTerm: "\<lbrakk>\<not> need_second_arg binop v1\<rbrakk> | |
| 161 | \<Longrightarrow> | |
| 162 | G\<turnstile>(\<langle>BinOp binop (Lit v1) e2\<rangle>,Norm s) | |
| 163 | \<mapsto>1 (\<langle>Lit v1\<rangle>,Norm s)" | |
| 164 | | BinOp: "G\<turnstile>(\<langle>BinOp binop (Lit v1) (Lit v2)\<rangle>,Norm s) | |
| 165 | \<mapsto>1 (\<langle>Lit (eval_binop binop v1 v2)\<rangle>,Norm s)" | |
| 13384 | 166 | (* Maybe its more convenient to add: need_second_arg as precondition to BinOp | 
| 167 | to make the choice between BinOpTerm and BinOp deterministic *) | |
| 168 | ||
| 21765 | 169 | | Super: "G\<turnstile>(\<langle>Super\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit (val_this s)\<rangle>,Norm s)" | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 170 | |
| 21765 | 171 | | AccVA: "\<lbrakk>G\<turnstile>(\<langle>va\<rangle>,Norm s) \<mapsto>1 (\<langle>va'\<rangle>,s') \<rbrakk> | 
| 172 | \<Longrightarrow> | |
| 173 | G\<turnstile>(\<langle>Acc va\<rangle>,Norm s) \<mapsto>1 (\<langle>Acc va'\<rangle>,s')" | |
| 174 | | Acc: "\<lbrakk>groundVar va; ((v,vf),s') = the_var G (Norm s) va\<rbrakk> | |
| 175 | \<Longrightarrow> | |
| 176 | G\<turnstile>(\<langle>Acc va\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit v\<rangle>,s')" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 177 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 178 | (* | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 179 | AccLVar: "G\<turnstile>(\<langle>Acc (LVar vn)\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit (fst (lvar vn s))\<rangle>,Norm s)" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 180 | AccFVar: "\<lbrakk>((v,vf),s') = fvar statDeclC stat fn a (Norm s)\<rbrakk> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 181 | \<Longrightarrow> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 182 |           G\<turnstile>(\<langle>Acc ({accC,statDeclC,stat}(Lit a)..fn)\<rangle>,Norm s) 
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 183 | \<mapsto>1 (\<langle>Lit v\<rangle>,s')" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 184 | AccAVar: "\<lbrakk>((v,vf),s') = avar G i a (Norm s)\<rbrakk> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 185 | \<Longrightarrow> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 186 | G\<turnstile>(\<langle>Acc ((Lit a).[Lit i])\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit v\<rangle>,s')" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 187 | *) | 
| 21765 | 188 | | AssVA: "\<lbrakk>G\<turnstile>(\<langle>va\<rangle>,Norm s) \<mapsto>1 (\<langle>va'\<rangle>,s')\<rbrakk> | 
| 189 | \<Longrightarrow> | |
| 190 | G\<turnstile>(\<langle>va:=e\<rangle>,Norm s) \<mapsto>1 (\<langle>va':=e\<rangle>,s')" | |
| 191 | | AssE: "\<lbrakk>groundVar va; G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> | |
| 192 | \<Longrightarrow> | |
| 193 | G\<turnstile>(\<langle>va:=e\<rangle>,Norm s) \<mapsto>1 (\<langle>va:=e'\<rangle>,s')" | |
| 194 | | Ass: "\<lbrakk>groundVar va; ((w,f),s') = the_var G (Norm s) va\<rbrakk> | |
| 195 | \<Longrightarrow> | |
| 196 | G\<turnstile>(\<langle>va:=(Lit v)\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit v\<rangle>,assign f v s')" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 197 | |
| 21765 | 198 | | CondC: "\<lbrakk>G\<turnstile>(\<langle>e0\<rangle>,Norm s) \<mapsto>1 (\<langle>e0'\<rangle>,s')\<rbrakk> | 
| 199 | \<Longrightarrow> | |
| 200 | G\<turnstile>(\<langle>e0? e1:e2\<rangle>,Norm s) \<mapsto>1 (\<langle>e0'? e1:e2\<rangle>,s')" | |
| 201 | | Cond: "G\<turnstile>(\<langle>Lit b? e1:e2\<rangle>,Norm s) \<mapsto>1 (\<langle>if the_Bool b then e1 else e2\<rangle>,Norm s)" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 202 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 203 | |
| 21765 | 204 | | CallTarget: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> | 
| 205 | \<Longrightarrow> | |
| 206 |                G\<turnstile>(\<langle>{accC,statT,mode}e\<cdot>mn({pTs}args)\<rangle>,Norm s) 
 | |
| 207 |                 \<mapsto>1 (\<langle>{accC,statT,mode}e'\<cdot>mn({pTs}args)\<rangle>,s')"
 | |
| 208 | | CallArgs: "\<lbrakk>G\<turnstile>(\<langle>args\<rangle>,Norm s) \<mapsto>1 (\<langle>args'\<rangle>,s')\<rbrakk> | |
| 209 | \<Longrightarrow> | |
| 210 |                G\<turnstile>(\<langle>{accC,statT,mode}Lit a\<cdot>mn({pTs}args)\<rangle>,Norm s) 
 | |
| 211 |                 \<mapsto>1 (\<langle>{accC,statT,mode}Lit a\<cdot>mn({pTs}args')\<rangle>,s')"
 | |
| 212 | | Call: "\<lbrakk>groundExprs args; vs = map the_val args; | |
| 213 | D = invocation_declclass G mode s a statT \<lparr>name=mn,parTs=pTs\<rparr>; | |
| 214 | s'=init_lvars G D \<lparr>name=mn,parTs=pTs\<rparr> mode a' vs (Norm s)\<rbrakk> | |
| 215 | \<Longrightarrow> | |
| 216 |                G\<turnstile>(\<langle>{accC,statT,mode}Lit a\<cdot>mn({pTs}args)\<rangle>,Norm s) 
 | |
| 217 | \<mapsto>1 (\<langle>Callee (locals s) (Methd D \<lparr>name=mn,parTs=pTs\<rparr>)\<rangle>,s')" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 218 | |
| 21765 | 219 | | Callee: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'::expr\<rangle>,s')\<rbrakk> | 
| 220 | \<Longrightarrow> | |
| 221 | G\<turnstile>(\<langle>Callee lcls_caller e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 222 | |
| 21765 | 223 | | CalleeRet: "G\<turnstile>(\<langle>Callee lcls_caller (Lit v)\<rangle>,Norm s) | 
| 224 | \<mapsto>1 (\<langle>Lit v\<rangle>,(set_lvars lcls_caller (Norm s)))" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 225 | |
| 21765 | 226 | | Methd: "G\<turnstile>(\<langle>Methd D sig\<rangle>,Norm s) \<mapsto>1 (\<langle>body G D sig\<rangle>,Norm s)" | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 227 | |
| 21765 | 228 | | Body: "G\<turnstile>(\<langle>Body D c\<rangle>,Norm s) \<mapsto>1 (\<langle>InsInitE (Init D) (Body D c)\<rangle>,Norm s)" | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 229 | |
| 21765 | 230 | | InsInitBody: | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 231 | "\<lbrakk>G\<turnstile>(\<langle>c\<rangle>,Norm s) \<mapsto>1 (\<langle>c'\<rangle>,s')\<rbrakk> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 232 | \<Longrightarrow> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 233 | G\<turnstile>(\<langle>InsInitE Skip (Body D c)\<rangle>,Norm s) \<mapsto>1(\<langle>InsInitE Skip (Body D c')\<rangle>,s')" | 
| 21765 | 234 | | InsInitBodyRet: | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 235 | "G\<turnstile>(\<langle>InsInitE Skip (Body D Skip)\<rangle>,Norm s) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 236 | \<mapsto>1 (\<langle>Lit (the ((locals s) Result))\<rangle>,abupd (absorb Ret) (Norm s))" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 237 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 238 | (* LVar: "G\<turnstile>(LVar vn,Norm s)" is already evaluated *) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 239 | |
| 21765 | 240 | | FVar: "\<lbrakk>\<not> inited statDeclC (globs s)\<rbrakk> | 
| 241 | \<Longrightarrow> | |
| 242 |          G\<turnstile>(\<langle>{accC,statDeclC,stat}e..fn\<rangle>,Norm s) 
 | |
| 243 |           \<mapsto>1 (\<langle>InsInitV (Init statDeclC) ({accC,statDeclC,stat}e..fn)\<rangle>,Norm s)"
 | |
| 244 | | InsInitFVarE: | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 245 | "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 246 | \<Longrightarrow> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 247 |        G\<turnstile>(\<langle>InsInitV Skip ({accC,statDeclC,stat}e..fn)\<rangle>,Norm s) 
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 248 |         \<mapsto>1 (\<langle>InsInitV Skip ({accC,statDeclC,stat}e'..fn)\<rangle>,s')"
 | 
| 21765 | 249 | | InsInitFVar: | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 250 |       "G\<turnstile>(\<langle>InsInitV Skip ({accC,statDeclC,stat}Lit a..fn)\<rangle>,Norm s) 
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 251 |         \<mapsto>1 (\<langle>{accC,statDeclC,stat}Lit a..fn\<rangle>,Norm s)"
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 252 | --  {* Notice, that we do not have literal values for @{text vars}. 
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 253 | The rules for accessing variables (@{text Acc}) and assigning to variables 
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 254 | (@{text Ass}), test this with the predicate @{text groundVar}.  After 
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 255 | initialisation is done and the @{text FVar} is evaluated, we can't just 
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 256 | throw away the @{text InsInitFVar} term and return a literal value, as in the 
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 257 | cases of @{text New}  or @{text NewC}. Instead we just return the evaluated 
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 258 | @{text FVar} and test for initialisation in the rule @{text FVar}. 
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 259 | *} | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 260 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 261 | |
| 21765 | 262 | | AVarE1: "\<lbrakk>G\<turnstile>(\<langle>e1\<rangle>,Norm s) \<mapsto>1 (\<langle>e1'\<rangle>,s')\<rbrakk> | 
| 263 | \<Longrightarrow> | |
| 264 | G\<turnstile>(\<langle>e1.[e2]\<rangle>,Norm s) \<mapsto>1 (\<langle>e1'.[e2]\<rangle>,s')" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 265 | |
| 21765 | 266 | | AVarE2: "G\<turnstile>(\<langle>e2\<rangle>,Norm s) \<mapsto>1 (\<langle>e2'\<rangle>,s') | 
| 267 | \<Longrightarrow> | |
| 268 | G\<turnstile>(\<langle>Lit a.[e2]\<rangle>,Norm s) \<mapsto>1 (\<langle>Lit a.[e2']\<rangle>,s')" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 269 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 270 | (* AVar: \<langle>(Lit a).[Lit i]\<rangle> is fully evaluated *) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 271 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 272 | (* evaluation of expression lists *) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 273 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 274 |   -- {* @{text Nil}  is fully evaluated *}
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 275 | |
| 21765 | 276 | | ConsHd: "\<lbrakk>G\<turnstile>(\<langle>e::expr\<rangle>,Norm s) \<mapsto>1 (\<langle>e'::expr\<rangle>,s')\<rbrakk> | 
| 277 | \<Longrightarrow> | |
| 278 | G\<turnstile>(\<langle>e#es\<rangle>,Norm s) \<mapsto>1 (\<langle>e'#es\<rangle>,s')" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 279 | |
| 21765 | 280 | | ConsTl: "\<lbrakk>G\<turnstile>(\<langle>es\<rangle>,Norm s) \<mapsto>1 (\<langle>es'\<rangle>,s')\<rbrakk> | 
| 281 | \<Longrightarrow> | |
| 282 | G\<turnstile>(\<langle>(Lit v)#es\<rangle>,Norm s) \<mapsto>1 (\<langle>(Lit v)#es'\<rangle>,s')" | |
| 12854 | 283 | |
| 284 | (* execution of statements *) | |
| 285 | ||
| 286 | (* cf. 14.5 *) | |
| 21765 | 287 | | Skip: "G\<turnstile>(\<langle>Skip\<rangle>,Norm s) \<mapsto>1 (\<langle>SKIP\<rangle>,Norm s)" | 
| 12854 | 288 | |
| 21765 | 289 | | ExprE: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> | 
| 290 | \<Longrightarrow> | |
| 291 | G\<turnstile>(\<langle>Expr e\<rangle>,Norm s) \<mapsto>1 (\<langle>Expr e'\<rangle>,s')" | |
| 292 | | Expr: "G\<turnstile>(\<langle>Expr (Lit v)\<rangle>,Norm s) \<mapsto>1 (\<langle>Skip\<rangle>,Norm s)" | |
| 12854 | 293 | |
| 294 | ||
| 21765 | 295 | | LabC: "\<lbrakk>G\<turnstile>(\<langle>c\<rangle>,Norm s) \<mapsto>1 (\<langle>c'\<rangle>,s')\<rbrakk> | 
| 296 | \<Longrightarrow> | |
| 297 | G\<turnstile>(\<langle>l\<bullet> c\<rangle>,Norm s) \<mapsto>1 (\<langle>l\<bullet> c'\<rangle>,s')" | |
| 298 | | Lab: "G\<turnstile>(\<langle>l\<bullet> Skip\<rangle>,s) \<mapsto>1 (\<langle>Skip\<rangle>, abupd (absorb l) s)" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 299 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 300 | (* cf. 14.2 *) | 
| 21765 | 301 | | CompC1: "\<lbrakk>G\<turnstile>(\<langle>c1\<rangle>,Norm s) \<mapsto>1 (\<langle>c1'\<rangle>,s')\<rbrakk> | 
| 302 | \<Longrightarrow> | |
| 303 | G\<turnstile>(\<langle>c1;; c2\<rangle>,Norm s) \<mapsto>1 (\<langle>c1';; c2\<rangle>,s')" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 304 | |
| 21765 | 305 | | Comp: "G\<turnstile>(\<langle>Skip;; c2\<rangle>,Norm s) \<mapsto>1 (\<langle>c2\<rangle>,Norm s)" | 
| 12854 | 306 | |
| 307 | (* cf. 14.8.2 *) | |
| 21765 | 308 | | IfE: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle> ,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> | 
| 309 | \<Longrightarrow> | |
| 310 | G\<turnstile>(\<langle>If(e) s1 Else s2\<rangle>,Norm s) \<mapsto>1 (\<langle>If(e') s1 Else s2\<rangle>,s')" | |
| 311 | | If: "G\<turnstile>(\<langle>If(Lit v) s1 Else s2\<rangle>,Norm s) | |
| 312 | \<mapsto>1 (\<langle>if the_Bool v then s1 else s2\<rangle>,Norm s)" | |
| 12854 | 313 | |
| 314 | (* cf. 14.10, 14.10.1 *) | |
| 21765 | 315 | | Loop: "G\<turnstile>(\<langle>l\<bullet> While(e) c\<rangle>,Norm s) | 
| 316 | \<mapsto>1 (\<langle>If(e) (Cont l\<bullet>c;; l\<bullet> While(e) c) Else Skip\<rangle>,Norm s)" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 317 | |
| 21765 | 318 | | Jmp: "G\<turnstile>(\<langle>Jmp j\<rangle>,Norm s) \<mapsto>1 (\<langle>Skip\<rangle>,(Some (Jump j), s))" | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 319 | |
| 21765 | 320 | | ThrowE: "\<lbrakk>G\<turnstile>(\<langle>e\<rangle>,Norm s) \<mapsto>1 (\<langle>e'\<rangle>,s')\<rbrakk> | 
| 321 | \<Longrightarrow> | |
| 322 | G\<turnstile>(\<langle>Throw e\<rangle>,Norm s) \<mapsto>1 (\<langle>Throw e'\<rangle>,s')" | |
| 323 | | Throw: "G\<turnstile>(\<langle>Throw (Lit a)\<rangle>,Norm s) \<mapsto>1 (\<langle>Skip\<rangle>,abupd (throw a) (Norm s))" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 324 | |
| 21765 | 325 | | TryC1: "\<lbrakk>G\<turnstile>(\<langle>c1\<rangle>,Norm s) \<mapsto>1 (\<langle>c1'\<rangle>,s')\<rbrakk> | 
| 326 | \<Longrightarrow> | |
| 327 | G\<turnstile>(\<langle>Try c1 Catch(C vn) c2\<rangle>, Norm s) \<mapsto>1 (\<langle>Try c1' Catch(C vn) c2\<rangle>,s')" | |
| 328 | | Try: "\<lbrakk>G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s'\<rbrakk> | |
| 329 | \<Longrightarrow> | |
| 330 | G\<turnstile>(\<langle>Try Skip Catch(C vn) c2\<rangle>, s) | |
| 331 | \<mapsto>1 (if G,s'\<turnstile>catch C then (\<langle>c2\<rangle>,new_xcpt_var vn s') | |
| 332 | else (\<langle>Skip\<rangle>,s'))" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 333 | |
| 21765 | 334 | | FinC1: "\<lbrakk>G\<turnstile>(\<langle>c1\<rangle>,Norm s) \<mapsto>1 (\<langle>c1'\<rangle>,s')\<rbrakk> | 
| 335 | \<Longrightarrow> | |
| 336 | G\<turnstile>(\<langle>c1 Finally c2\<rangle>,Norm s) \<mapsto>1 (\<langle>c1' Finally c2\<rangle>,s')" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 337 | |
| 21765 | 338 | | Fin: "G\<turnstile>(\<langle>Skip Finally c2\<rangle>,(a,s)) \<mapsto>1 (\<langle>FinA a c2\<rangle>,Norm s)" | 
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 339 | |
| 21765 | 340 | | FinAC: "\<lbrakk>G\<turnstile>(\<langle>c\<rangle>,s) \<mapsto>1 (\<langle>c'\<rangle>,s')\<rbrakk> | 
| 341 | \<Longrightarrow> | |
| 342 | G\<turnstile>(\<langle>FinA a c\<rangle>,s) \<mapsto>1 (\<langle>FinA a c'\<rangle>,s')" | |
| 343 | | FinA: "G\<turnstile>(\<langle>FinA a Skip\<rangle>,s) \<mapsto>1 (\<langle>Skip\<rangle>,abupd (abrupt_if (a\<noteq>None) a) s)" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 344 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 345 | |
| 21765 | 346 | | Init1: "\<lbrakk>inited C (globs s)\<rbrakk> | 
| 347 | \<Longrightarrow> | |
| 348 | G\<turnstile>(\<langle>Init C\<rangle>,Norm s) \<mapsto>1 (\<langle>Skip\<rangle>,Norm s)" | |
| 349 | | Init: "\<lbrakk>the (class G C)=c; \<not> inited C (globs s)\<rbrakk> | |
| 350 | \<Longrightarrow> | |
| 351 | G\<turnstile>(\<langle>Init C\<rangle>,Norm s) | |
| 352 | \<mapsto>1 (\<langle>(if C = Object then Skip else (Init (super c)));; | |
| 353 | Expr (Callee (locals s) (InsInitE (init c) SKIP))\<rangle> | |
| 354 | ,Norm (init_class_obj G C s))" | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 355 | -- {* @{text InsInitE} is just used as trick to embed the statement 
 | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 356 | @{text "init c"} into an expression*} 
 | 
| 21765 | 357 | | InsInitESKIP: | 
| 358 | "G\<turnstile>(\<langle>InsInitE Skip SKIP\<rangle>,Norm s) \<mapsto>1 (\<langle>SKIP\<rangle>,Norm s)" | |
| 359 | ||
| 360 | abbreviation | |
| 361 |   stepn:: "[prog, term \<times> state,nat,term \<times> state] \<Rightarrow> bool" ("_\<turnstile>_ \<mapsto>_ _"[61,82,82] 81)
 | |
| 30952 
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
 haftmann parents: 
28524diff
changeset | 362 |   where "G\<turnstile>p \<mapsto>n p' \<equiv> (p,p') \<in> {(x, y). step G x y}^^n"
 | 
| 21765 | 363 | |
| 364 | abbreviation | |
| 365 |   steptr:: "[prog,term \<times> state,term \<times> state] \<Rightarrow> bool" ("_\<turnstile>_ \<mapsto>* _"[61,82,82] 81)
 | |
| 366 |   where "G\<turnstile>p \<mapsto>* p' \<equiv> (p,p') \<in> {(x, y). step G x y}\<^sup>*"
 | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 367 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 368 | (* Equivalenzen: | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 369 | Bigstep zu Smallstep komplett. | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 370 | Smallstep zu Bigstep, nur wenn nicht die Ausdrücke Callee, FinA ,\<dots> | 
| 12925 
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
 schirmer parents: 
12858diff
changeset | 371 | *) | 
| 12854 | 372 | |
| 13337 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 373 | (* | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 374 | lemma imp_eval_trans: | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 375 | assumes eval: "G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v,s1)" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 376 | shows trans: "G\<turnstile>(t,s0) \<mapsto>* (\<langle>Lit v\<rangle>,s1)" | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 377 | *) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 378 | (* Jetzt muss man bei trans natürlich wieder unterscheiden: Stmt, Expr, Var! | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 379 | Sowas blödes: | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 380 | Am besten den Terminus ground auf Var,Stmt,Expr hochziehen und dann | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 381 | the_vals definieren\<dots> | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 382 | G\<turnstile>(t,s0) \<mapsto>* (t',s1) \<and> the_vals t' = v | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 383 | *) | 
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 384 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 385 | |
| 
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
 schirmer parents: 
12925diff
changeset | 386 | end |