author | clasohm |
Wed, 02 Nov 1994 11:50:09 +0100 | |
changeset 156 | fd1be45b64bf |
parent 126 | 872f866e630f |
child 171 | 16c4ea954511 |
permissions | -rw-r--r-- |
0 | 1 |
(* Title: Substitutions/uterm.ML |
2 |
Author: Martin Coen, Cambridge University Computer Laboratory |
|
3 |
Copyright 1993 University of Cambridge |
|
4 |
||
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
5 |
Simple term structure for unifiation. |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
6 |
Binary trees with leaves that are constants or variables. |
0 | 7 |
*) |
8 |
||
9 |
open UTerm; |
|
10 |
||
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
11 |
val uterm_con_defs = [VAR_def, CONST_def, COMB_def]; |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
12 |
|
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
13 |
goal UTerm.thy "uterm(A) = A <+> A <+> (uterm(A) <*> uterm(A))"; |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
14 |
let val rew = rewrite_rule uterm_con_defs in |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
15 |
by (fast_tac (univ_cs addSIs (equalityI :: map rew uterm.intrs) |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
16 |
addEs [rew uterm.elim]) 1) |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
17 |
end; |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
18 |
val uterm_unfold = result(); |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
19 |
|
0 | 20 |
(** the uterm functional **) |
21 |
||
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
22 |
(*This justifies using uterm in other recursive type definitions*) |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
23 |
goalw UTerm.thy uterm.defs "!!A B. A<=B ==> uterm(A) <= uterm(B)"; |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
24 |
by (REPEAT (ares_tac (lfp_mono::basic_monos) 1)); |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
25 |
val uterm_mono = result(); |
0 | 26 |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
27 |
(** Type checking rules -- uterm creates well-founded sets **) |
0 | 28 |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
29 |
goalw UTerm.thy (uterm_con_defs @ uterm.defs) "uterm(sexp) <= sexp"; |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
30 |
by (rtac lfp_lowerbound 1); |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
31 |
by (fast_tac (univ_cs addIs sexp.intrs@[sexp_In0I,sexp_In1I]) 1); |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
32 |
val uterm_sexp = result(); |
0 | 33 |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
34 |
(* A <= sexp ==> uterm(A) <= sexp *) |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
35 |
val uterm_subset_sexp = standard ([uterm_mono, uterm_sexp] MRS subset_trans); |
0 | 36 |
|
37 |
(** Induction **) |
|
38 |
||
39 |
(*Induction for the type 'a uterm *) |
|
40 |
val prems = goalw UTerm.thy [Var_def,Const_def,Comb_def] |
|
41 |
"[| !!x.P(Var(x)); !!x.P(Const(x)); \ |
|
42 |
\ !!u v. [| P(u); P(v) |] ==> P(Comb(u,v)) |] ==> P(t)"; |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
43 |
by (rtac (Rep_uterm_inverse RS subst) 1); (*types force good instantiation*) |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
44 |
by (rtac (Rep_uterm RS uterm.induct) 1); |
0 | 45 |
by (REPEAT (ares_tac prems 1 |
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
46 |
ORELSE eresolve_tac [rangeE, ssubst, Abs_uterm_inverse RS subst] 1)); |
0 | 47 |
val uterm_induct = result(); |
48 |
||
49 |
(*Perform induction on xs. *) |
|
50 |
fun uterm_ind_tac a M = |
|
51 |
EVERY [res_inst_tac [("t",a)] uterm_induct M, |
|
52 |
rename_last_tac a ["1"] (M+1)]; |
|
53 |
||
54 |
||
55 |
(*** Isomorphisms ***) |
|
56 |
||
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
57 |
goal UTerm.thy "inj(Rep_uterm)"; |
0 | 58 |
by (rtac inj_inverseI 1); |
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
59 |
by (rtac Rep_uterm_inverse 1); |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
60 |
val inj_Rep_uterm = result(); |
0 | 61 |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
62 |
goal UTerm.thy "inj_onto(Abs_uterm,uterm(range(Leaf)))"; |
0 | 63 |
by (rtac inj_onto_inverseI 1); |
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
64 |
by (etac Abs_uterm_inverse 1); |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
65 |
val inj_onto_Abs_uterm = result(); |
0 | 66 |
|
67 |
(** Distinctness of constructors **) |
|
68 |
||
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
69 |
goalw UTerm.thy uterm_con_defs "~ CONST(c) = COMB(u,v)"; |
0 | 70 |
by (rtac notI 1); |
71 |
by (etac (In1_inject RS (In0_not_In1 RS notE)) 1); |
|
72 |
val CONST_not_COMB = result(); |
|
73 |
val COMB_not_CONST = standard (CONST_not_COMB RS not_sym); |
|
74 |
val CONST_neq_COMB = standard (CONST_not_COMB RS notE); |
|
75 |
val COMB_neq_CONST = sym RS CONST_neq_COMB; |
|
76 |
||
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
77 |
goalw UTerm.thy uterm_con_defs "~ COMB(u,v) = VAR(x)"; |
0 | 78 |
by (rtac In1_not_In0 1); |
79 |
val COMB_not_VAR = result(); |
|
80 |
val VAR_not_COMB = standard (COMB_not_VAR RS not_sym); |
|
81 |
val COMB_neq_VAR = standard (COMB_not_VAR RS notE); |
|
82 |
val VAR_neq_COMB = sym RS COMB_neq_VAR; |
|
83 |
||
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
84 |
goalw UTerm.thy uterm_con_defs "~ VAR(x) = CONST(c)"; |
0 | 85 |
by (rtac In0_not_In1 1); |
86 |
val VAR_not_CONST = result(); |
|
87 |
val CONST_not_VAR = standard (VAR_not_CONST RS not_sym); |
|
88 |
val VAR_neq_CONST = standard (VAR_not_CONST RS notE); |
|
89 |
val CONST_neq_VAR = sym RS VAR_neq_CONST; |
|
90 |
||
91 |
||
92 |
goalw UTerm.thy [Const_def,Comb_def] "~ Const(c) = Comb(u,v)"; |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
93 |
by (rtac (CONST_not_COMB RS (inj_onto_Abs_uterm RS inj_onto_contraD)) 1); |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
94 |
by (REPEAT (resolve_tac (uterm.intrs @ [rangeI, Rep_uterm]) 1)); |
0 | 95 |
val Const_not_Comb = result(); |
96 |
val Comb_not_Const = standard (Const_not_Comb RS not_sym); |
|
97 |
val Const_neq_Comb = standard (Const_not_Comb RS notE); |
|
98 |
val Comb_neq_Const = sym RS Const_neq_Comb; |
|
99 |
||
100 |
goalw UTerm.thy [Comb_def,Var_def] "~ Comb(u,v) = Var(x)"; |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
101 |
by (rtac (COMB_not_VAR RS (inj_onto_Abs_uterm RS inj_onto_contraD)) 1); |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
102 |
by (REPEAT (resolve_tac (uterm.intrs @ [rangeI, Rep_uterm]) 1)); |
0 | 103 |
val Comb_not_Var = result(); |
104 |
val Var_not_Comb = standard (Comb_not_Var RS not_sym); |
|
105 |
val Comb_neq_Var = standard (Comb_not_Var RS notE); |
|
106 |
val Var_neq_Comb = sym RS Comb_neq_Var; |
|
107 |
||
108 |
goalw UTerm.thy [Var_def,Const_def] "~ Var(x) = Const(c)"; |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
109 |
by (rtac (VAR_not_CONST RS (inj_onto_Abs_uterm RS inj_onto_contraD)) 1); |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
110 |
by (REPEAT (resolve_tac (uterm.intrs @ [rangeI, Rep_uterm]) 1)); |
0 | 111 |
val Var_not_Const = result(); |
112 |
val Const_not_Var = standard (Var_not_Const RS not_sym); |
|
113 |
val Var_neq_Const = standard (Var_not_Const RS notE); |
|
114 |
val Const_neq_Var = sym RS Var_neq_Const; |
|
115 |
||
116 |
||
117 |
(** Injectiveness of CONST and Const **) |
|
118 |
||
119 |
val inject_cs = HOL_cs addSEs [Scons_inject] |
|
120 |
addSDs [In0_inject,In1_inject]; |
|
121 |
||
122 |
goalw UTerm.thy [VAR_def] "(VAR(M)=VAR(N)) = (M=N)"; |
|
123 |
by (fast_tac inject_cs 1); |
|
124 |
val VAR_VAR_eq = result(); |
|
125 |
||
126 |
goalw UTerm.thy [CONST_def] "(CONST(M)=CONST(N)) = (M=N)"; |
|
127 |
by (fast_tac inject_cs 1); |
|
128 |
val CONST_CONST_eq = result(); |
|
129 |
||
130 |
goalw UTerm.thy [COMB_def] "(COMB(K,L)=COMB(M,N)) = (K=M & L=N)"; |
|
131 |
by (fast_tac inject_cs 1); |
|
132 |
val COMB_COMB_eq = result(); |
|
133 |
||
134 |
val VAR_inject = standard (VAR_VAR_eq RS iffD1); |
|
135 |
val CONST_inject = standard (CONST_CONST_eq RS iffD1); |
|
136 |
val COMB_inject = standard (COMB_COMB_eq RS iffD1 RS conjE); |
|
137 |
||
138 |
||
139 |
(*For reasoning about abstract uterm constructors*) |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
140 |
val uterm_cs = set_cs addIs uterm.intrs @ [Rep_uterm] |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
141 |
addSEs [CONST_neq_COMB,COMB_neq_VAR,VAR_neq_CONST, |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
142 |
COMB_neq_CONST,VAR_neq_COMB,CONST_neq_VAR, |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
143 |
COMB_inject] |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
144 |
addSDs [VAR_inject,CONST_inject, |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
145 |
inj_onto_Abs_uterm RS inj_ontoD, |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
146 |
inj_Rep_uterm RS injD, Leaf_inject]; |
0 | 147 |
|
148 |
goalw UTerm.thy [Var_def] "(Var(x)=Var(y)) = (x=y)"; |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
149 |
by (fast_tac uterm_cs 1); |
0 | 150 |
val Var_Var_eq = result(); |
151 |
val Var_inject = standard (Var_Var_eq RS iffD1); |
|
152 |
||
153 |
goalw UTerm.thy [Const_def] "(Const(x)=Const(y)) = (x=y)"; |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
154 |
by (fast_tac uterm_cs 1); |
0 | 155 |
val Const_Const_eq = result(); |
156 |
val Const_inject = standard (Const_Const_eq RS iffD1); |
|
157 |
||
158 |
goalw UTerm.thy [Comb_def] "(Comb(u,v)=Comb(x,y)) = (u=x & v=y)"; |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
159 |
by (fast_tac uterm_cs 1); |
0 | 160 |
val Comb_Comb_eq = result(); |
161 |
val Comb_inject = standard (Comb_Comb_eq RS iffD1 RS conjE); |
|
162 |
||
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
163 |
val [major] = goal UTerm.thy "VAR(M): uterm(A) ==> M : A"; |
0 | 164 |
by (rtac (major RS setup_induction) 1); |
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
165 |
by (etac uterm.induct 1); |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
166 |
by (ALLGOALS (fast_tac uterm_cs)); |
0 | 167 |
val VAR_D = result(); |
168 |
||
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
169 |
val [major] = goal UTerm.thy "CONST(M): uterm(A) ==> M : A"; |
0 | 170 |
by (rtac (major RS setup_induction) 1); |
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
171 |
by (etac uterm.induct 1); |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
172 |
by (ALLGOALS (fast_tac uterm_cs)); |
0 | 173 |
val CONST_D = result(); |
174 |
||
175 |
val [major] = goal UTerm.thy |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
176 |
"COMB(M,N): uterm(A) ==> M: uterm(A) & N: uterm(A)"; |
0 | 177 |
by (rtac (major RS setup_induction) 1); |
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
178 |
by (etac uterm.induct 1); |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
179 |
by (ALLGOALS (fast_tac uterm_cs)); |
0 | 180 |
val COMB_D = result(); |
181 |
||
182 |
(*Basic ss with constructors and their freeness*) |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
183 |
val uterm_free_simps = uterm.intrs @ |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
184 |
[Const_not_Comb,Comb_not_Var,Var_not_Const, |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
185 |
Comb_not_Const,Var_not_Comb,Const_not_Var, |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
186 |
Var_Var_eq,Const_Const_eq,Comb_Comb_eq, |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
187 |
CONST_not_COMB,COMB_not_VAR,VAR_not_CONST, |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
188 |
COMB_not_CONST,VAR_not_COMB,CONST_not_VAR, |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
189 |
VAR_VAR_eq,CONST_CONST_eq,COMB_COMB_eq]; |
0 | 190 |
val uterm_free_ss = HOL_ss addsimps uterm_free_simps; |
191 |
||
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
192 |
goal UTerm.thy "!u. t~=Comb(t,u)"; |
0 | 193 |
by (uterm_ind_tac "t" 1); |
194 |
by (rtac (Var_not_Comb RS allI) 1); |
|
195 |
by (rtac (Const_not_Comb RS allI) 1); |
|
196 |
by (asm_simp_tac uterm_free_ss 1); |
|
197 |
val t_not_Comb_t = result(); |
|
198 |
||
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
199 |
goal UTerm.thy "!t. u~=Comb(t,u)"; |
0 | 200 |
by (uterm_ind_tac "u" 1); |
201 |
by (rtac (Var_not_Comb RS allI) 1); |
|
202 |
by (rtac (Const_not_Comb RS allI) 1); |
|
203 |
by (asm_simp_tac uterm_free_ss 1); |
|
204 |
val u_not_Comb_u = result(); |
|
205 |
||
206 |
||
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
207 |
(*** UTerm_rec -- by wf recursion on pred_sexp ***) |
0 | 208 |
|
209 |
val UTerm_rec_unfold = |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
210 |
[UTerm_rec_def, wf_pred_sexp RS wf_trancl] MRS def_wfrec; |
0 | 211 |
|
212 |
(** conversion rules **) |
|
213 |
||
214 |
goalw UTerm.thy [VAR_def] "UTerm_rec(VAR(x),b,c,d) = b(x)"; |
|
215 |
by (rtac (UTerm_rec_unfold RS trans) 1); |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
216 |
by (simp_tac (HOL_ss addsimps [Case_In0]) 1); |
0 | 217 |
val UTerm_rec_VAR = result(); |
218 |
||
219 |
goalw UTerm.thy [CONST_def] "UTerm_rec(CONST(x),b,c,d) = c(x)"; |
|
220 |
by (rtac (UTerm_rec_unfold RS trans) 1); |
|
221 |
by (simp_tac (HOL_ss addsimps [Case_In0,Case_In1]) 1); |
|
222 |
val UTerm_rec_CONST = result(); |
|
223 |
||
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
224 |
goalw UTerm.thy [COMB_def] |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
225 |
"!!M N. [| M: sexp; N: sexp |] ==> \ |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
226 |
\ UTerm_rec(COMB(M,N), b, c, d) = \ |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
227 |
\ d(M, N, UTerm_rec(M,b,c,d), UTerm_rec(N,b,c,d))"; |
0 | 228 |
by (rtac (UTerm_rec_unfold RS trans) 1); |
229 |
by (simp_tac (HOL_ss addsimps [Split,Case_In1]) 1); |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
230 |
by (asm_simp_tac (pred_sexp_ss addsimps [In1_def]) 1); |
0 | 231 |
val UTerm_rec_COMB = result(); |
232 |
||
233 |
(*** uterm_rec -- by UTerm_rec ***) |
|
234 |
||
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
235 |
val Rep_uterm_in_sexp = |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
236 |
Rep_uterm RS (range_Leaf_subset_sexp RS uterm_subset_sexp RS subsetD); |
0 | 237 |
|
126
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
238 |
val uterm_rec_simps = |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
239 |
uterm.intrs @ |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
240 |
[UTerm_rec_VAR, UTerm_rec_CONST, UTerm_rec_COMB, |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
241 |
Abs_uterm_inverse, Rep_uterm_inverse, |
872f866e630f
Subst/UTerm: updated for inductive defs; streamlined proofs
lcp
parents:
0
diff
changeset
|
242 |
Rep_uterm, rangeI, inj_Leaf, Inv_f_f, Rep_uterm_in_sexp]; |
0 | 243 |
val uterm_rec_ss = HOL_ss addsimps uterm_rec_simps; |
244 |
||
245 |
goalw UTerm.thy [uterm_rec_def, Var_def] "uterm_rec(Var(x),b,c,d) = b(x)"; |
|
246 |
by (simp_tac uterm_rec_ss 1); |
|
247 |
val uterm_rec_Var = result(); |
|
248 |
||
249 |
goalw UTerm.thy [uterm_rec_def, Const_def] "uterm_rec(Const(x),b,c,d) = c(x)"; |
|
250 |
by (simp_tac uterm_rec_ss 1); |
|
251 |
val uterm_rec_Const = result(); |
|
252 |
||
253 |
goalw UTerm.thy [uterm_rec_def, Comb_def] |
|
254 |
"uterm_rec(Comb(u,v),b,c,d) = d(u,v,uterm_rec(u,b,c,d),uterm_rec(v,b,c,d))"; |
|
255 |
by (simp_tac uterm_rec_ss 1); |
|
256 |
val uterm_rec_Comb = result(); |
|
257 |
||
258 |
val uterm_simps = [UTerm_rec_VAR, UTerm_rec_CONST, UTerm_rec_COMB, |
|
259 |
uterm_rec_Var, uterm_rec_Const, uterm_rec_Comb]; |
|
260 |
val uterm_ss = uterm_free_ss addsimps uterm_simps; |
|
261 |
||
262 |
||
263 |
(**********) |
|
264 |
||
265 |
val uterm_rews = [uterm_rec_Var,uterm_rec_Const,uterm_rec_Comb, |
|
266 |
t_not_Comb_t,u_not_Comb_u, |
|
267 |
Const_not_Comb,Comb_not_Var,Var_not_Const, |
|
268 |
Comb_not_Const,Var_not_Comb,Const_not_Var, |
|
269 |
Var_Var_eq,Const_Const_eq,Comb_Comb_eq]; |
|
270 |