|
1461
|
1 |
(* Title: ZF/IMP/Denotation.ML
|
|
482
|
2 |
ID: $Id$
|
|
1461
|
3 |
Author: Heiko Loetzbeyer & Robert Sandner, TUM
|
|
482
|
4 |
Copyright 1994 TUM
|
|
|
5 |
*)
|
|
|
6 |
|
|
|
7 |
open Denotation;
|
|
|
8 |
|
|
4298
|
9 |
(** Rewrite Rules for A,B,C **)
|
|
|
10 |
Addsimps [A_nat_def,A_loc_def,A_op1_def,A_op2_def];
|
|
|
11 |
Addsimps [B_true_def,B_false_def,B_op_def,B_not_def,B_and_def,B_or_def];
|
|
|
12 |
Addsimps [C_skip_def,C_assign_def,C_comp_def,C_if_def,C_while_def];
|
|
482
|
13 |
|
|
4298
|
14 |
(** Type_intr for A **)
|
|
482
|
15 |
|
|
|
16 |
val A_type = prove_goal Denotation.thy
|
|
1461
|
17 |
"!!a.[|a:aexp; sigma:loc->nat|] ==> A(a,sigma):nat"
|
|
511
|
18 |
(fn _ => [(etac aexp.induct 1),
|
|
4298
|
19 |
(ALLGOALS Asm_simp_tac),
|
|
4091
|
20 |
(ALLGOALS (fast_tac (claset() addSIs [apply_type])))]);
|
|
482
|
21 |
|
|
4298
|
22 |
(** Type_intr for B **)
|
|
482
|
23 |
|
|
|
24 |
val B_type = prove_goal Denotation.thy
|
|
1461
|
25 |
"!!b. [|b:bexp; sigma:loc->nat|] ==> B(b,sigma):bool"
|
|
511
|
26 |
(fn _ => [(etac bexp.induct 1),
|
|
4298
|
27 |
(ALLGOALS Asm_simp_tac),
|
|
4091
|
28 |
(ALLGOALS (fast_tac (claset()
|
|
482
|
29 |
addSIs [apply_type,A_type]@bool_typechecks)))]);
|
|
|
30 |
|
|
4298
|
31 |
(** C_subset **)
|
|
482
|
32 |
|
|
|
33 |
val C_subset = prove_goal Denotation.thy
|
|
1461
|
34 |
"!!c. c:com ==> C(c) <= (loc->nat)*(loc->nat)"
|
|
511
|
35 |
(fn _ => [(etac com.induct 1),
|
|
4298
|
36 |
(ALLGOALS Asm_simp_tac),
|
|
4091
|
37 |
(ALLGOALS (fast_tac (claset() addDs [lfp_subset RS subsetD])))]);
|
|
482
|
38 |
|
|
4298
|
39 |
(** Type_elims for C **)
|
|
482
|
40 |
|
|
|
41 |
val C_type = prove_goal Denotation.thy
|
|
1461
|
42 |
"[| <x,y>:C(c); c:com; \
|
|
|
43 |
\ !!c. [| x:loc->nat; y:loc->nat |] ==> R |] \
|
|
|
44 |
\ ==> R"
|
|
482
|
45 |
(fn prems => [(cut_facts_tac prems 1),
|
|
4091
|
46 |
(fast_tac (claset() addSIs prems
|
|
482
|
47 |
addDs [(C_subset RS subsetD)]) 1)]);
|
|
|
48 |
|
|
|
49 |
val C_type_fst = prove_goal Denotation.thy
|
|
1461
|
50 |
"[| x:C(c); c:com; \
|
|
|
51 |
\ !!c. [| fst(x):loc->nat |] ==> R |] \
|
|
|
52 |
\ ==> R"
|
|
482
|
53 |
(fn prems => [(cut_facts_tac prems 1),
|
|
|
54 |
(resolve_tac prems 1),
|
|
|
55 |
(dtac (C_subset RS subsetD) 1),
|
|
|
56 |
(atac 1),
|
|
|
57 |
(etac SigmaE 1),
|
|
2469
|
58 |
(Asm_simp_tac 1)]);
|
|
482
|
59 |
|
|
|
60 |
|
|
4298
|
61 |
(** bnd_mono (nat->nat*nat->nat,Gamma(b,c) **)
|
|
482
|
62 |
|
|
|
63 |
val Gamma_bnd_mono = prove_goalw Denotation.thy [bnd_mono_def,Gamma_def]
|
|
1461
|
64 |
"!!c. c:com ==> bnd_mono ((loc->nat)*(loc->nat),Gamma(b,c))"
|
|
4091
|
65 |
(fn prems => [(best_tac (claset() addEs [C_type]) 1)]);
|
|
482
|
66 |
|
|
4298
|
67 |
(** End ***)
|