| author | paulson | 
| Tue, 01 May 2001 17:16:32 +0200 | |
| changeset 11276 | f8353c722d4e | 
| parent 11070 | cc421547e744 | 
| child 11372 | 648795477bb5 | 
| permissions | -rw-r--r-- | 
| 8011 | 1 | (* Title: HOL/MicroJava/J/Eval.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: David von Oheimb | |
| 4 | Copyright 1999 Technische Universitaet Muenchen | |
| 11070 | 5 | *) | 
| 8011 | 6 | |
| 11070 | 7 | header "Operational Evaluation (big-step) Semantics" | 
| 8011 | 8 | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 9 | theory Eval = State + WellType: | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 10 | |
| 8011 | 11 | consts | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 12 | eval :: "java_mb prog => (xstate \<times> expr \<times> val \<times> xstate) set" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 13 | evals :: "java_mb prog => (xstate \<times> expr list \<times> val list \<times> xstate) set" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 14 | exec :: "java_mb prog => (xstate \<times> stmt \<times> xstate) set" | 
| 8011 | 15 | |
| 16 | syntax | |
| 10056 | 17 | eval :: "[java_mb prog,xstate,expr,val,xstate] => bool " | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 18 |           ("_ \<turnstile> _ -_\<succ>_-> _" [51,82,60,82,82] 81)
 | 
| 8082 | 19 | evals:: "[java_mb prog,xstate,expr list, | 
| 10056 | 20 | val list,xstate] => bool " | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 21 |           ("_ \<turnstile> _ -_[\<succ>]_-> _" [51,82,60,51,82] 81)
 | 
| 10056 | 22 | exec :: "[java_mb prog,xstate,stmt, xstate] => bool " | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 23 |           ("_ \<turnstile> _ -_-> _" [51,82,60,82] 81)
 | 
| 10061 
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
 kleing parents: 
10056diff
changeset | 24 | |
| 
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
 kleing parents: 
10056diff
changeset | 25 | syntax (HTML) | 
| 
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
 kleing parents: 
10056diff
changeset | 26 | eval :: "[java_mb prog,xstate,expr,val,xstate] => bool " | 
| 10828 
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
 oheimb parents: 
10763diff
changeset | 27 |           ("_ |- _ -_>_-> _" [51,82,60,82,82] 81)
 | 
| 10061 
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
 kleing parents: 
10056diff
changeset | 28 | evals:: "[java_mb prog,xstate,expr list, | 
| 
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
 kleing parents: 
10056diff
changeset | 29 | val list,xstate] => bool " | 
| 10828 
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
 oheimb parents: 
10763diff
changeset | 30 |           ("_ |- _ -_[>]_-> _" [51,82,60,51,82] 81)
 | 
| 10061 
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
 kleing parents: 
10056diff
changeset | 31 | exec :: "[java_mb prog,xstate,stmt, xstate] => bool " | 
| 10828 
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
 oheimb parents: 
10763diff
changeset | 32 |           ("_ |- _ -_-> _" [51,82,60,82] 81)
 | 
| 10061 
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
 kleing parents: 
10056diff
changeset | 33 | |
| 8011 | 34 | |
| 35 | translations | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 36 | "G\<turnstile>s -e \<succ> v-> (x,s')" <= "(s, e, v, x, s') \<in> eval G" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 37 | "G\<turnstile>s -e \<succ> v-> s' " == "(s, e, v, s') \<in> eval G" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 38 | "G\<turnstile>s -e[\<succ>]v-> (x,s')" <= "(s, e, v, x, s') \<in> evals G" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 39 | "G\<turnstile>s -e[\<succ>]v-> s' " == "(s, e, v, s') \<in> evals G" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 40 | "G\<turnstile>s -c -> (x,s')" <= "(s, c, x, s') \<in> exec G" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 41 | "G\<turnstile>s -c -> s' " == "(s, c, s') \<in> exec G" | 
| 8011 | 42 | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 43 | inductive "eval G" "evals G" "exec G" intros | 
| 8011 | 44 | |
| 45 | (* evaluation of expressions *) | |
| 46 | ||
| 47 | (* cf. 15.5 *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 48 | XcptE:"G\<turnstile>(Some xc,s) -e\<succ>arbitrary-> (Some xc,s)" | 
| 8011 | 49 | |
| 50 | (* cf. 15.8.1 *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 51 | NewC: "[| h = heap s; (a,x) = new_Addr h; | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 52 | h'= h(a\<mapsto>(C,init_vars (fields (G,C)))) |] ==> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 53 | G\<turnstile>Norm s -NewC C\<succ>Addr a-> c_hupd h' (x,s)" | 
| 8011 | 54 | |
| 55 | (* cf. 15.15 *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 56 | Cast: "[| G\<turnstile>Norm s0 -e\<succ>v-> (x1,s1); | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 57 | x2 = raise_if (\<not> cast_ok G C (heap s1) v) ClassCast x1 |] ==> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 58 | G\<turnstile>Norm s0 -Cast C e\<succ>v-> (x2,s1)" | 
| 8011 | 59 | |
| 60 | (* cf. 15.7.1 *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 61 | Lit: "G\<turnstile>Norm s -Lit v\<succ>v-> Norm s" | 
| 8011 | 62 | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 63 | BinOp:"[| G\<turnstile>Norm s -e1\<succ>v1-> s1; | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 64 | G\<turnstile>s1 -e2\<succ>v2-> s2; | 
| 10056 | 65 | v = (case bop of Eq => Bool (v1 = v2) | 
| 66 | | Add => Intg (the_Intg v1 + the_Intg v2)) |] ==> | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 67 | G\<turnstile>Norm s -BinOp bop e1 e2\<succ>v-> s2" | 
| 9240 | 68 | |
| 8011 | 69 | (* cf. 15.13.1, 15.2 *) | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 70 | LAcc: "G\<turnstile>Norm s -LAcc v\<succ>the (locals s v)-> Norm s" | 
| 8011 | 71 | |
| 72 | (* cf. 15.25.1 *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 73 | LAss: "[| G\<turnstile>Norm s -e\<succ>v-> (x,(h,l)); | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 74 | l' = (if x = None then l(va\<mapsto>v) else l) |] ==> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 75 | G\<turnstile>Norm s -va::=e\<succ>v-> (x,(h,l'))" | 
| 8011 | 76 | |
| 77 | ||
| 78 | (* cf. 15.10.1, 15.2 *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 79 | FAcc: "[| G\<turnstile>Norm s0 -e\<succ>a'-> (x1,s1); | 
| 10056 | 80 | v = the (snd (the (heap s1 (the_Addr a'))) (fn,T)) |] ==> | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 81 |          G\<turnstile>Norm s0 -{T}e..fn\<succ>v-> (np a' x1,s1)"
 | 
| 8011 | 82 | |
| 83 | (* cf. 15.25.1 *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 84 | FAss: "[| G\<turnstile> Norm s0 -e1\<succ>a'-> (x1,s1); a = the_Addr a'; | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 85 | G\<turnstile>(np a' x1,s1) -e2\<succ>v -> (x2,s2); | 
| 10056 | 86 | h = heap s2; (c,fs) = the (h a); | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 87 | h' = h(a\<mapsto>(c,(fs((fn,T)\<mapsto>v)))) |] ==> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 88 |          G\<turnstile>Norm s0 -{T}e1..fn:=e2\<succ>v-> c_hupd h' (x2,s2)"
 | 
| 8011 | 89 | |
| 90 | (* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5, 14.15 *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 91 | Call: "[| G\<turnstile>Norm s0 -e\<succ>a'-> s1; a = the_Addr a'; | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 92 | G\<turnstile>s1 -ps[\<succ>]pvs-> (x,(h,l)); dynT = fst (the (h a)); | 
| 10056 | 93 | (md,rT,pns,lvars,blk,res) = the (method (G,dynT) (mn,pTs)); | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 94 | G\<turnstile>(np a' x,(h,(init_vars lvars)(pns[\<mapsto>]pvs)(This\<mapsto>a'))) -blk-> s3; | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 95 | G\<turnstile> s3 -res\<succ>v -> (x4,s4) |] ==> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 96 |          G\<turnstile>Norm s0 -{C}e..mn({pTs}ps)\<succ>v-> (x4,(heap s4,l))"
 | 
| 8011 | 97 | |
| 98 | ||
| 99 | (* evaluation of expression lists *) | |
| 100 | ||
| 101 | (* cf. 15.5 *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 102 | XcptEs:"G\<turnstile>(Some xc,s) -e[\<succ>]arbitrary-> (Some xc,s)" | 
| 8011 | 103 | |
| 104 | (* cf. 15.11.??? *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 105 | Nil: "G\<turnstile>Norm s0 -[][\<succ>][]-> Norm s0" | 
| 8011 | 106 | |
| 107 | (* cf. 15.6.4 *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 108 | Cons: "[| G\<turnstile>Norm s0 -e \<succ> v -> s1; | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 109 | G\<turnstile> s1 -es[\<succ>]vs-> s2 |] ==> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 110 | G\<turnstile>Norm s0 -e#es[\<succ>]v#vs-> s2" | 
| 8011 | 111 | |
| 112 | (* execution of statements *) | |
| 113 | ||
| 114 | (* cf. 14.1 *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 115 | XcptS:"G\<turnstile>(Some xc,s) -c-> (Some xc,s)" | 
| 8011 | 116 | |
| 117 | (* cf. 14.5 *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 118 | Skip: "G\<turnstile>Norm s -Skip-> Norm s" | 
| 8011 | 119 | |
| 120 | (* cf. 14.7 *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 121 | Expr: "[| G\<turnstile>Norm s0 -e\<succ>v-> s1 |] ==> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 122 | G\<turnstile>Norm s0 -Expr e-> s1" | 
| 8011 | 123 | |
| 124 | (* cf. 14.2 *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 125 | Comp: "[| G\<turnstile>Norm s0 -c1-> s1; | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 126 | G\<turnstile> s1 -c2-> s2|] ==> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 127 | G\<turnstile>Norm s0 -c1;; c2-> s2" | 
| 8011 | 128 | |
| 129 | (* cf. 14.8.2 *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 130 | Cond: "[| G\<turnstile>Norm s0 -e\<succ>v-> s1; | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 131 | G\<turnstile> s1 -(if the_Bool v then c1 else c2)-> s2|] ==> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 132 | G\<turnstile>Norm s0 -If(e) c1 Else c2-> s2" | 
| 8011 | 133 | |
| 134 | (* cf. 14.10, 14.10.1 *) | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 135 | LoopF:"[| G\<turnstile>Norm s0 -e\<succ>v-> s1; \<not>the_Bool v |] ==> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 136 | G\<turnstile>Norm s0 -While(e) c-> s1" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 137 | LoopT:"[| G\<turnstile>Norm s0 -e\<succ>v-> s1; the_Bool v; | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 138 | G\<turnstile>s1 -c-> s2; G\<turnstile>s2 -While(e) c-> s3 |] ==> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 139 | G\<turnstile>Norm s0 -While(e) c-> s3" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 140 | |
| 11040 | 141 | lemmas eval_evals_exec_induct = eval_evals_exec.induct [split_format (complete)] | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 142 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 143 | lemma NewCI: "[|new_Addr (heap s) = (a,x); | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 144 | s' = c_hupd (heap s(a\<mapsto>(C,init_vars (fields (G,C))))) (x,s)|] ==> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 145 | G\<turnstile>Norm s -NewC C\<succ>Addr a-> s'" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 146 | apply (simp (no_asm_simp)) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 147 | apply (rule eval_evals_exec.NewC) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 148 | apply auto | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 149 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 150 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 151 | lemma eval_evals_exec_no_xcpt: | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 152 | "!!s s'. (G\<turnstile>(x,s) -e \<succ> v -> (x',s') --> x'=None --> x=None) \<and> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 153 | (G\<turnstile>(x,s) -es[\<succ>]vs-> (x',s') --> x'=None --> x=None) \<and> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 154 | (G\<turnstile>(x,s) -c -> (x',s') --> x'=None --> x=None)" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 155 | apply(simp (no_asm_simp) only: split_tupled_all) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 156 | apply(rule eval_evals_exec_induct) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 157 | apply(unfold c_hupd_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 158 | apply(simp_all) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 159 | done | 
| 10828 
b207d6d1bedc
improved evaluation judgment syntax; modified Loop rule
 oheimb parents: 
10763diff
changeset | 160 | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 161 | lemma eval_no_xcpt: "G\<turnstile>(x,s) -e\<succ>v-> (None,s') ==> x=None" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 162 | apply (drule eval_evals_exec_no_xcpt [THEN conjunct1, THEN mp]) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 163 | apply (fast) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 164 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 165 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 166 | lemma evals_no_xcpt: "G\<turnstile>(x,s) -e[\<succ>]v-> (None,s') ==> x=None" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 167 | apply (drule eval_evals_exec_no_xcpt [THEN conjunct2, THEN conjunct1, THEN mp]) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 168 | apply (fast) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 169 | done | 
| 8011 | 170 | |
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 171 | lemma eval_evals_exec_xcpt: | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 172 | "!!s s'. (G\<turnstile>(x,s) -e \<succ> v -> (x',s') --> x=Some xc --> x'=Some xc \<and> s'=s) \<and> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 173 | (G\<turnstile>(x,s) -es[\<succ>]vs-> (x',s') --> x=Some xc --> x'=Some xc \<and> s'=s) \<and> | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 174 | (G\<turnstile>(x,s) -c -> (x',s') --> x=Some xc --> x'=Some xc \<and> s'=s)" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 175 | apply (simp (no_asm_simp) only: split_tupled_all) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 176 | apply (rule eval_evals_exec_induct) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 177 | apply (unfold c_hupd_def) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 178 | apply (simp_all) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 179 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 180 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 181 | lemma eval_xcpt: "G\<turnstile>(Some xc,s) -e\<succ>v-> (x',s') ==> x'=Some xc \<and> s'=s" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 182 | apply (drule eval_evals_exec_xcpt [THEN conjunct1, THEN mp]) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 183 | apply (fast) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 184 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 185 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 186 | lemma exec_xcpt: "G\<turnstile>(Some xc,s) -s0-> (x',s') ==> x'=Some xc \<and> s'=s" | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 187 | apply (drule eval_evals_exec_xcpt [THEN conjunct2, THEN conjunct2, THEN mp]) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 188 | apply (fast) | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 189 | done | 
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 190 | |
| 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10927diff
changeset | 191 | end |