author | oheimb |
Fri, 14 Jul 2000 20:47:11 +0200 | |
changeset 9348 | f495dba0cb07 |
parent 9346 | 297dcbf64526 |
child 10042 | 7164dc0d24d8 |
permissions | -rw-r--r-- |
8011 | 1 |
(* Title: HOL/MicroJava/J/Eval.ML |
2 |
ID: $Id$ |
|
3 |
Author: David von Oheimb |
|
4 |
Copyright 1999 Technische Universitaet Muenchen |
|
5 |
*) |
|
6 |
||
9346
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
8011
diff
changeset
|
7 |
Goal "\\<lbrakk>new_Addr (heap s) = (a,x); \ |
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
8011
diff
changeset
|
8 |
\ s' = c_hupd (heap s(a\\<mapsto>(C,init_vars (fields (G,C))))) (x,s)\\<rbrakk> \\<Longrightarrow> \ |
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
8011
diff
changeset
|
9 |
\ G\\<turnstile>Norm s -NewC C\\<succ>Addr a\\<rightarrow> s'"; |
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
8011
diff
changeset
|
10 |
by (hyp_subst_tac 1); |
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
8011
diff
changeset
|
11 |
br eval_evals_exec.NewC 1; |
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
8011
diff
changeset
|
12 |
by Auto_tac; |
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
8011
diff
changeset
|
13 |
qed "NewCI"; |
297dcbf64526
re-structuring MicroJava; added Example; corrected := syntax; simplfied cast
oheimb
parents:
8011
diff
changeset
|
14 |
|
8011 | 15 |
Goal "\\<And>s s'. (G\\<turnstile>(x,s) -e \\<succ> v \\<rightarrow> (x',s') \\<longrightarrow> x'=None \\<longrightarrow> x=None) \\<and> \ |
16 |
\ (G\\<turnstile>(x,s) -es[\\<succ>]vs\\<rightarrow> (x',s') \\<longrightarrow> x'=None \\<longrightarrow> x=None) \\<and> \ |
|
17 |
\ (G\\<turnstile>(x,s) -c \\<rightarrow> (x',s') \\<longrightarrow> x'=None \\<longrightarrow> x=None)"; |
|
18 |
by(split_all_tac 1); |
|
19 |
by(rtac eval_evals_exec.induct 1); |
|
20 |
by(rewtac c_hupd_def); |
|
21 |
by(ALLGOALS Asm_full_simp_tac); |
|
22 |
qed "eval_evals_exec_no_xcpt"; |
|
23 |
||
24 |
val eval_no_xcpt = prove_goal thy "\\<And>X. G\\<turnstile>(x,s) -e\\<succ>v\\<rightarrow> (None,s') \\<Longrightarrow> x=None" (K [ |
|
25 |
dtac (eval_evals_exec_no_xcpt RS conjunct1 RS mp) 1, |
|
26 |
Fast_tac 1]); |
|
27 |
val evals_no_xcpt = prove_goal thy "\\<And>X. G\\<turnstile>(x,s) -e[\\<succ>]v\\<rightarrow> (None,s') \\<Longrightarrow> x=None" (K [ |
|
28 |
dtac (eval_evals_exec_no_xcpt RS conjunct2 RS conjunct1 RS mp) 1, |
|
29 |
Fast_tac 1]); |
|
30 |
||
31 |
val eval_evals_exec_xcpt = prove_goal thy |
|
32 |
"\\<And>s s'. (G\\<turnstile>(x,s) -e \\<succ> v \\<rightarrow> (x',s') \\<longrightarrow> x=Some xc \\<longrightarrow> x'=Some xc \\<and> s'=s) \\<and> \ |
|
33 |
\ (G\\<turnstile>(x,s) -es[\\<succ>]vs\\<rightarrow> (x',s') \\<longrightarrow> x=Some xc \\<longrightarrow> x'=Some xc \\<and> s'=s) \\<and> \ |
|
34 |
\ (G\\<turnstile>(x,s) -c \\<rightarrow> (x',s') \\<longrightarrow> x=Some xc \\<longrightarrow> x'=Some xc \\<and> s'=s)" |
|
35 |
(K [ |
|
36 |
split_all_tac 1, |
|
37 |
rtac eval_evals_exec.induct 1, |
|
38 |
rewtac c_hupd_def, |
|
39 |
ALLGOALS Asm_full_simp_tac]); |
|
40 |
val eval_xcpt = prove_goal thy |
|
41 |
"\\<And>X. G\\<turnstile>(Some xc,s) -e\\<succ>v\\<rightarrow> (x',s') \\<Longrightarrow> x'=Some xc \\<and> s'=s" (K [ |
|
42 |
dtac (eval_evals_exec_xcpt RS conjunct1 RS mp) 1, |
|
43 |
Fast_tac 1]); |
|
44 |
val exec_xcpt = prove_goal thy |
|
45 |
"\\<And>X. G\\<turnstile>(Some xc,s) -s0\\<rightarrow> (x',s') \\<Longrightarrow> x'=Some xc \\<and> s'=s" (K [ |
|
46 |
dtac (eval_evals_exec_xcpt RS conjunct2 RS conjunct2 RS mp) 1, |
|
47 |
Fast_tac 1]); |