1048
|
1 |
(* Title: Reduction.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Ole Rasmussen
|
|
4 |
Copyright 1995 University of Cambridge
|
|
5 |
Logic Image: ZF
|
|
6 |
*)
|
|
7 |
|
|
8 |
open Reduction;
|
|
9 |
|
|
10 |
|
|
11 |
(* ------------------------------------------------------------------------- *)
|
|
12 |
(* Setting up rulelists for reduction *)
|
|
13 |
(* ------------------------------------------------------------------------- *)
|
|
14 |
|
|
15 |
val red1D1 = Sred1.dom_subset RS subsetD RS SigmaD1;
|
|
16 |
val red1D2 = Sred1.dom_subset RS subsetD RS SigmaD2;
|
|
17 |
val redD1 = Sred.dom_subset RS subsetD RS SigmaD1;
|
|
18 |
val redD2 = Sred.dom_subset RS subsetD RS SigmaD2;
|
|
19 |
val par_red1D1 = Spar_red1.dom_subset RS subsetD RS SigmaD1;
|
|
20 |
val par_red1D2 = Spar_red1.dom_subset RS subsetD RS SigmaD2;
|
|
21 |
val par_redD1 = Spar_red.dom_subset RS subsetD RS SigmaD1;
|
|
22 |
val par_redD2 = Spar_red.dom_subset RS subsetD RS SigmaD2;
|
|
23 |
|
|
24 |
|
|
25 |
val reduc_cs = res_cs
|
|
26 |
addIs (Sred1.intrs@[Sred.one_step,Sred.refl]@Spar_red1.intrs@
|
|
27 |
[Spar_red.one_step,lambda.dom_subset RS subsetD,
|
|
28 |
unmark_type]@lambda.intrs@bool_typechecks)
|
|
29 |
addSEs [Spar_red1.mk_cases redexes.con_defs "Fun(t) =1=> Fun(u)"];
|
|
30 |
|
|
31 |
val reduc_ss = term_ss addsimps
|
|
32 |
(Sred1.intrs@[Sred.one_step,Sred.refl]@Spar_red1.intrs@
|
|
33 |
[Spar_red.one_step,substL_type,redD1,redD2,par_redD1,
|
|
34 |
par_redD2,par_red1D2,unmark_type]);
|
|
35 |
|
|
36 |
val reducL_ss = reduc_ss setloop (SELECT_GOAL (safe_tac reduc_cs));
|
|
37 |
|
|
38 |
val red1_induct = Sred1.mutual_induct RS spec RS spec RSN (2,rev_mp);
|
|
39 |
val red_induct = Sred.mutual_induct RS spec RS spec RSN (2,rev_mp);
|
|
40 |
val par_red1_induct = Spar_red1.mutual_induct RS spec RS spec RSN (2,rev_mp);
|
|
41 |
val par_red_induct = Spar_red.mutual_induct RS spec RS spec RSN (2,rev_mp);
|
|
42 |
|
|
43 |
(* ------------------------------------------------------------------------- *)
|
|
44 |
(* Lemmas for reduction *)
|
|
45 |
(* ------------------------------------------------------------------------- *)
|
|
46 |
|
|
47 |
goal Reduction.thy "!!u. m--->n ==> Fun(m) ---> Fun(n)";
|
|
48 |
by (eresolve_tac [red_induct] 1);
|
|
49 |
by (resolve_tac [Sred.trans] 3);
|
|
50 |
by (ALLGOALS (asm_simp_tac reduc_ss ));
|
|
51 |
val red_Fun = result();
|
|
52 |
|
|
53 |
goal Reduction.thy
|
|
54 |
"!!u.[|n:lambda; m ---> m'|]==> Apl(m,n)--->Apl(m',n)";
|
|
55 |
by (eresolve_tac [red_induct] 1);
|
|
56 |
by (resolve_tac [Sred.trans] 3);
|
|
57 |
by (ALLGOALS (asm_simp_tac reduc_ss ));
|
|
58 |
val red_Apll = result();
|
|
59 |
|
|
60 |
goal Reduction.thy
|
|
61 |
"!!u.[|n:lambda; m ---> m'|]==> Apl(n,m)--->Apl(n,m')";
|
|
62 |
by (eresolve_tac [red_induct] 1);
|
|
63 |
by (resolve_tac [Sred.trans] 3);
|
|
64 |
by (ALLGOALS (asm_simp_tac reduc_ss ));
|
|
65 |
val red_Aplr = result();
|
|
66 |
|
|
67 |
goal Reduction.thy
|
|
68 |
"!!u.[|m ---> m'; n--->n'|]==> Apl(m,n)--->Apl(m',n')";
|
|
69 |
by (res_inst_tac [("n","Apl(m',n)")] Sred.trans 1);
|
|
70 |
by (ALLGOALS (asm_simp_tac (reduc_ss addsimps [red_Apll,red_Aplr]) ));
|
|
71 |
val red_Apl = result();
|
|
72 |
|
|
73 |
goal Reduction.thy
|
|
74 |
"!!u.[|m:lambda; m':lambda; n:lambda; n':lambda; m ---> m'; n--->n'|]==> \
|
|
75 |
\ Apl(Fun(m),n)---> n'/m'";
|
|
76 |
by (res_inst_tac [("n","Apl(Fun(m'),n')")] Sred.trans 1);
|
|
77 |
by (ALLGOALS (asm_simp_tac (reduc_ss addsimps [red_Apl,red_Fun]) ));
|
|
78 |
val red_beta = result();
|
|
79 |
|
|
80 |
|
|
81 |
(* ------------------------------------------------------------------------- *)
|
|
82 |
(* Lemmas for parallel reduction *)
|
|
83 |
(* ------------------------------------------------------------------------- *)
|
|
84 |
|
|
85 |
|
|
86 |
goal Reduction.thy "!!u.m:lambda==> m =1=> m";
|
|
87 |
by (eresolve_tac [lambda.induct] 1);
|
|
88 |
by (ALLGOALS (asm_simp_tac reduc_ss ));
|
|
89 |
val refl_par_red1 = result();
|
|
90 |
|
|
91 |
goal Reduction.thy "!!u.m-1->n ==> m=1=>n";
|
|
92 |
by (eresolve_tac [red1_induct] 1);
|
|
93 |
by (ALLGOALS (asm_simp_tac (reduc_ss addsimps [refl_par_red1]) ));
|
|
94 |
val red1_par_red1 = result();
|
|
95 |
|
|
96 |
goal Reduction.thy "!!u.m--->n ==> m===>n";
|
|
97 |
by (eresolve_tac [red_induct] 1);
|
|
98 |
by (resolve_tac [Spar_red.trans] 3);
|
|
99 |
by (ALLGOALS (asm_simp_tac (reduc_ss addsimps [refl_par_red1,red1_par_red1]) ));
|
|
100 |
val red_par_red = result();
|
|
101 |
|
|
102 |
goal Reduction.thy "!!u.m===>n ==> m--->n";
|
|
103 |
by (eresolve_tac [par_red_induct] 1);
|
|
104 |
by (eresolve_tac [par_red1_induct] 1);
|
|
105 |
by (resolve_tac [Sred.trans] 5);
|
|
106 |
by (ALLGOALS (asm_simp_tac (reduc_ss addsimps [red_Fun,red_beta,red_Apl]) ));
|
|
107 |
val par_red_red = result();
|
|
108 |
|
|
109 |
|
|
110 |
(* ------------------------------------------------------------------------- *)
|
|
111 |
(* Simulation *)
|
|
112 |
(* ------------------------------------------------------------------------- *)
|
|
113 |
|
|
114 |
goal Reduction.thy
|
|
115 |
"!!u.m=1=>n ==> EX v. m|>v = n & m~v & regular(v)";
|
|
116 |
by (eresolve_tac [par_red1_induct] 1);
|
|
117 |
by (step_tac ZF_cs 1);
|
|
118 |
by (ALLGOALS (REPEAT o (resolve_tac [exI,conjI])));
|
|
119 |
by (TRYALL(resolve_tac [res_redex,res_App,res_Fun,res_Var]));
|
|
120 |
by (ALLGOALS (asm_simp_tac (reduc_ss)));
|
|
121 |
val simulation = result();
|
|
122 |
|
|
123 |
|
|
124 |
(* ------------------------------------------------------------------------- *)
|
|
125 |
(* commuting of unmark and subst *)
|
|
126 |
(* ------------------------------------------------------------------------- *)
|
|
127 |
|
|
128 |
goal Reduction.thy
|
|
129 |
"!!u.u:redexes ==> \
|
|
130 |
\ ALL k:nat. unmark(lift_rec(u,k)) = lift_rec(unmark(u),k)";
|
|
131 |
by (eresolve_tac [redexes.induct] 1);
|
|
132 |
by (ALLGOALS (asm_full_simp_tac (addsplit reduc_ss)));
|
|
133 |
val unmmark_lift_rec = result();
|
|
134 |
|
|
135 |
goal Reduction.thy
|
|
136 |
"!!u.v:redexes ==> ALL k:nat.ALL u:redexes. \
|
|
137 |
\ unmark(subst_rec(u,v,k)) = subst_rec(unmark(u),unmark(v),k)";
|
|
138 |
by (eresolve_tac [redexes.induct] 1);
|
|
139 |
by (ALLGOALS (asm_full_simp_tac
|
|
140 |
((addsplit reduc_ss) addsimps [unmmark_lift_rec])));
|
|
141 |
val unmmark_subst_rec = result();
|
|
142 |
|
|
143 |
|
|
144 |
(* ------------------------------------------------------------------------- *)
|
|
145 |
(* Completeness *)
|
|
146 |
(* ------------------------------------------------------------------------- *)
|
|
147 |
|
|
148 |
goal Reduction.thy
|
|
149 |
"!!u.u~v ==> regular(v) --> unmark(u) =1=> unmark(u|>v)";
|
|
150 |
by (eresolve_tac [comp_induct] 1);
|
|
151 |
by (ALLGOALS (asm_full_simp_tac (reducL_ss addsimps [unmmark_subst_rec]) ));
|
|
152 |
by (dres_inst_tac [("psi", "Fun(?u) =1=> ?w")] asm_rl 1);
|
|
153 |
by (asm_full_simp_tac reducL_ss 1);
|
|
154 |
val completeness_l = result();
|
|
155 |
|
|
156 |
goal Reduction.thy
|
|
157 |
"!!u.[|u:lambda; u~v; regular(v)|]==> u =1=> unmark(u|>v)";
|
|
158 |
by (dres_inst_tac [("u1","u")] (completeness_l RS mp) 1);
|
|
159 |
by (ALLGOALS (asm_full_simp_tac (reduc_ss addsimps [lambda_unmark]) ));
|
|
160 |
val completeness = result();
|