author | wenzelm |
Mon, 29 Nov 1999 15:52:49 +0100 | |
changeset 8039 | a901bafe4578 |
parent 6154 | 6a00a5baef2b |
child 11318 | 6536fb8c9fc6 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/Coind/Values.ML |
916 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Jacob Frost, Cambridge University Computer Laboratory |
916 | 4 |
Copyright 1995 University of Cambridge |
5 |
*) |
|
6 |
||
7 |
open Values; |
|
8 |
||
9 |
(* Elimination rules *) |
|
10 |
||
11 |
val prems = goalw Values.thy (Part_def::(tl (tl Val_ValEnv.defs))) |
|
12 |
"[| ve:ValEnv; !!m.[| ve=ve_mk(m); m:PMap(ExVar,Val) |] ==> Q |] ==> Q"; |
|
13 |
by (cut_facts_tac prems 1); |
|
14 |
by (etac CollectE 1); |
|
15 |
by (etac exE 1); |
|
16 |
by (etac Val_ValEnv.elim 1); |
|
17 |
by (eresolve_tac prems 3); |
|
18 |
by (rewrite_tac Val_ValEnv.con_defs); |
|
2469 | 19 |
by (ALLGOALS Fast_tac); |
916 | 20 |
qed "ValEnvE"; |
21 |
||
22 |
val prems = goalw Values.thy (Part_def::[hd (tl Val_ValEnv.defs)]) |
|
23 |
"[| v:Val; \ |
|
24 |
\ !!c. [| v = v_const(c); c:Const |] ==> Q;\ |
|
25 |
\ !!e ve x. \ |
|
26 |
\ [| v = v_clos(x,e,ve); x:ExVar; e:Exp; ve:ValEnv |] ==> Q \ |
|
27 |
\ |] ==> \ |
|
28 |
\ Q"; |
|
29 |
by (cut_facts_tac prems 1); |
|
30 |
by (etac CollectE 1); |
|
31 |
by (etac exE 1); |
|
32 |
by (etac Val_ValEnv.elim 1); |
|
33 |
by (eresolve_tac prems 1); |
|
34 |
by (assume_tac 1); |
|
35 |
by (eresolve_tac prems 1); |
|
36 |
by (assume_tac 1); |
|
37 |
by (assume_tac 1); |
|
38 |
by (assume_tac 1); |
|
39 |
by (rewrite_tac Val_ValEnv.con_defs); |
|
2469 | 40 |
by (Fast_tac 1); |
916 | 41 |
qed "ValE"; |
42 |
||
43 |
(* Nonempty sets *) |
|
44 |
||
6154
6a00a5baef2b
automatic insertion of datatype intr rules into claset
paulson
parents:
6068
diff
changeset
|
45 |
Goalw (QPair_def::QInl_def::QInr_def::Val_ValEnv.con_defs) |
916 | 46 |
"v_clos(x,e,ve) ~= 0"; |
6154
6a00a5baef2b
automatic insertion of datatype intr rules into claset
paulson
parents:
6068
diff
changeset
|
47 |
by (Blast_tac 1); |
916 | 48 |
qed "v_closNE"; |
49 |
||
6154
6a00a5baef2b
automatic insertion of datatype intr rules into claset
paulson
parents:
6068
diff
changeset
|
50 |
Addsimps [v_closNE]; |
6a00a5baef2b
automatic insertion of datatype intr rules into claset
paulson
parents:
6068
diff
changeset
|
51 |
AddSEs [v_closNE RS notE]; |
6a00a5baef2b
automatic insertion of datatype intr rules into claset
paulson
parents:
6068
diff
changeset
|
52 |
|
6a00a5baef2b
automatic insertion of datatype intr rules into claset
paulson
parents:
6068
diff
changeset
|
53 |
Goalw (QPair_def::QInl_def::QInr_def::Val_ValEnv.con_defs) |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
54 |
"c:Const ==> v_const(c) ~= 0"; |
916 | 55 |
by (dtac constNEE 1); |
6154
6a00a5baef2b
automatic insertion of datatype intr rules into claset
paulson
parents:
6068
diff
changeset
|
56 |
by Auto_tac; |
916 | 57 |
qed "v_constNE"; |
58 |
||
6154
6a00a5baef2b
automatic insertion of datatype intr rules into claset
paulson
parents:
6068
diff
changeset
|
59 |
Addsimps [v_constNE]; |
6a00a5baef2b
automatic insertion of datatype intr rules into claset
paulson
parents:
6068
diff
changeset
|
60 |
|
916 | 61 |
(* Proving that the empty set is not a value *) |
62 |
||
5137 | 63 |
Goal "v:Val ==> v ~= 0"; |
916 | 64 |
by (etac ValE 1); |
6154
6a00a5baef2b
automatic insertion of datatype intr rules into claset
paulson
parents:
6068
diff
changeset
|
65 |
by Auto_tac; |
916 | 66 |
qed "ValNEE"; |
67 |
||
68 |
(* Equalities for value environments *) |
|
69 |
||
6046 | 70 |
Goal "[| ve:ValEnv; v ~=0 |] ==> ve_dom(ve_owr(ve,x,v)) = ve_dom(ve) Un {x}"; |
916 | 71 |
by (etac ValEnvE 1); |
6154
6a00a5baef2b
automatic insertion of datatype intr rules into claset
paulson
parents:
6068
diff
changeset
|
72 |
by (auto_tac (claset(), |
6a00a5baef2b
automatic insertion of datatype intr rules into claset
paulson
parents:
6068
diff
changeset
|
73 |
simpset() addsimps [map_domain_owr])); |
916 | 74 |
qed "ve_dom_owr"; |
75 |
||
6068 | 76 |
Goal "ve:ValEnv \ |
77 |
\ ==> ve_app(ve_owr(ve,y,v),x) = (if x=y then v else ve_app(ve,x))"; |
|
916 | 78 |
by (etac ValEnvE 1); |
6046 | 79 |
by (asm_simp_tac (simpset() addsimps [map_app_owr]) 1); |
80 |
qed "ve_app_owr"; |
|
81 |
||
82 |
Goal "ve:ValEnv ==> ve_app(ve_owr(ve,x,v),x) = v"; |
|
83 |
by (etac ValEnvE 1); |
|
84 |
by (Asm_full_simp_tac 1); |
|
916 | 85 |
by (rtac map_app_owr1 1); |
86 |
qed "ve_app_owr1"; |
|
87 |
||
6046 | 88 |
Goal "ve:ValEnv ==> x ~= y ==> ve_app(ve_owr(ve,x,v),y) = ve_app(ve,y)"; |
916 | 89 |
by (etac ValEnvE 1); |
6046 | 90 |
by (Asm_full_simp_tac 1); |
916 | 91 |
by (rtac map_app_owr2 1); |
2469 | 92 |
by (Fast_tac 1); |
916 | 93 |
qed "ve_app_owr2"; |
94 |
||
95 |
(* Introduction rules for operators on value environments *) |
|
96 |
||
6046 | 97 |
Goal "[| ve:ValEnv; x:ve_dom(ve) |] ==> ve_app(ve,x):Val"; |
916 | 98 |
by (etac ValEnvE 1); |
99 |
by (hyp_subst_tac 1); |
|
6046 | 100 |
by (Asm_full_simp_tac 1); |
916 | 101 |
by (rtac pmap_appI 1); |
102 |
by (assume_tac 1); |
|
103 |
by (assume_tac 1); |
|
104 |
qed "ve_appI"; |
|
105 |
||
6046 | 106 |
Goal "[| ve:ValEnv; x:ve_dom(ve) |] ==> x:ExVar"; |
916 | 107 |
by (etac ValEnvE 1); |
108 |
by (hyp_subst_tac 1); |
|
6046 | 109 |
by (Asm_full_simp_tac 1); |
916 | 110 |
by (rtac pmap_domainD 1); |
111 |
by (assume_tac 1); |
|
112 |
by (assume_tac 1); |
|
113 |
qed "ve_domI"; |
|
114 |
||
5068 | 115 |
Goalw [ve_emp_def] "ve_emp:ValEnv"; |
916 | 116 |
by (rtac Val_ValEnv.ve_mk_I 1); |
117 |
by (rtac pmap_empI 1); |
|
118 |
qed "ve_empI"; |
|
119 |
||
6046 | 120 |
Goal "[|ve:ValEnv; x:ExVar; v:Val |] ==> ve_owr(ve,x,v):ValEnv"; |
916 | 121 |
by (etac ValEnvE 1); |
122 |
by (hyp_subst_tac 1); |
|
6046 | 123 |
by (Asm_full_simp_tac 1); |
916 | 124 |
by (rtac Val_ValEnv.ve_mk_I 1); |
125 |
by (etac pmap_owrI 1); |
|
126 |
by (assume_tac 1); |
|
127 |
by (assume_tac 1); |
|
128 |
qed "ve_owrI"; |
|
129 |
||
130 |
||
131 |