author | wenzelm |
Wed, 31 May 2000 14:30:28 +0200 | |
changeset 9011 | 0cfc347f8d19 |
parent 8817 | 1c48b3732399 |
child 9444 | 13b10be222bf |
permissions | -rw-r--r-- |
3302 | 1 |
(* Title: TFL/post |
2 |
ID: $Id$ |
|
3 |
Author: Konrad Slind, Cambridge University Computer Laboratory |
|
4 |
Copyright 1997 University of Cambridge |
|
5 |
||
6 |
Postprocessing of TFL definitions |
|
7 |
*) |
|
8 |
||
3191 | 9 |
signature TFL = |
10 |
sig |
|
7262 | 11 |
|
12 |
val trace : bool ref |
|
13 |
||
2112 | 14 |
structure Prim : TFL_sig |
6524 | 15 |
val quiet_mode : bool ref |
16 |
val message : string -> unit |
|
2112 | 17 |
|
3191 | 18 |
val tgoalw : theory -> thm list -> thm list -> thm list |
19 |
val tgoal: theory -> thm list -> thm list |
|
2112 | 20 |
|
3405 | 21 |
val std_postprocessor : simpset -> theory |
2112 | 22 |
-> {induction:thm, rules:thm, TCs:term list list} |
23 |
-> {induction:thm, rules:thm, nested_tcs:thm list} |
|
24 |
||
7262 | 25 |
val define_i : theory -> xstring -> term |
26 |
-> simpset * thm list (*allows special simplication*) |
|
27 |
-> term list |
|
8622
870a58dd0ddd
the simplification rules returned from TFL are now paired with the row they
nipkow
parents:
8526
diff
changeset
|
28 |
-> theory * {rules:(thm*int)list, induct:thm, tcs:term list} |
3331 | 29 |
|
7262 | 30 |
val define : theory -> xstring -> string |
31 |
-> simpset * thm list |
|
32 |
-> string list |
|
8622
870a58dd0ddd
the simplification rules returned from TFL are now paired with the row they
nipkow
parents:
8526
diff
changeset
|
33 |
-> theory * {rules:(thm*int)list, induct:thm, tcs:term list} |
2112 | 34 |
|
7262 | 35 |
val defer_i : theory -> xstring |
6498 | 36 |
-> thm list (* congruence rules *) |
37 |
-> term list -> theory * thm |
|
38 |
||
7262 | 39 |
val defer : theory -> xstring |
6498 | 40 |
-> thm list |
41 |
-> string list -> theory * thm |
|
8631 | 42 |
|
3191 | 43 |
end; |
44 |
||
45 |
||
46 |
structure Tfl: TFL = |
|
2112 | 47 |
struct |
48 |
structure Prim = Prim |
|
3391
5e45dd3b64e9
More de-HOLification: using Free, Const, etc. instead of mk_var, mk_const
paulson
parents:
3331
diff
changeset
|
49 |
structure S = USyntax |
2112 | 50 |
|
8817 | 51 |
fun read_term thy = Sign.simple_read_term (Theory.sign_of thy) HOLogic.termT; |
52 |
||
53 |
||
6524 | 54 |
(* messages *) |
55 |
||
56 |
val quiet_mode = ref false; |
|
57 |
fun message s = if ! quiet_mode then () else writeln s; |
|
58 |
||
6498 | 59 |
val trace = Prim.trace |
60 |
||
8817 | 61 |
|
6498 | 62 |
(*--------------------------------------------------------------------------- |
63 |
* Extract termination goals so that they can be put it into a goalstack, or |
|
64 |
* have a tactic directly applied to them. |
|
65 |
*--------------------------------------------------------------------------*) |
|
66 |
fun termination_goals rules = |
|
67 |
map (#1 o Type.freeze_thaw o HOLogic.dest_Trueprop) |
|
68 |
(foldr (fn (th,A) => union_term (prems_of th, A)) (rules, [])); |
|
3191 | 69 |
|
6498 | 70 |
(*--------------------------------------------------------------------------- |
71 |
* Finds the termination conditions in (highly massaged) definition and |
|
72 |
* puts them into a goalstack. |
|
73 |
*--------------------------------------------------------------------------*) |
|
74 |
fun tgoalw thy defs rules = |
|
75 |
case termination_goals rules of |
|
76 |
[] => error "tgoalw: no termination conditions to prove" |
|
77 |
| L => goalw_cterm defs |
|
78 |
(Thm.cterm_of (Theory.sign_of thy) |
|
79 |
(HOLogic.mk_Trueprop(USyntax.list_mk_conj L))); |
|
80 |
||
81 |
fun tgoal thy = tgoalw thy []; |
|
82 |
||
83 |
(*--------------------------------------------------------------------------- |
|
84 |
* Three postprocessors are applied to the definition. It |
|
85 |
* attempts to prove wellfoundedness of the given relation, simplifies the |
|
86 |
* non-proved termination conditions, and finally attempts to prove the |
|
87 |
* simplified termination conditions. |
|
3405 | 88 |
*--------------------------------------------------------------------------*) |
6498 | 89 |
fun std_postprocessor ss = |
90 |
Prim.postprocess |
|
91 |
{WFtac = REPEAT (ares_tac [wf_empty, wf_pred_nat, |
|
92 |
wf_measure, wf_inv_image, |
|
93 |
wf_lex_prod, wf_less_than, wf_trancl] 1), |
|
94 |
terminator = asm_simp_tac ss 1 |
|
8526
0be2c98f15a7
replaced best_tac by force_tac, allowing some arithmetic reasoning
paulson
parents:
7262
diff
changeset
|
95 |
THEN TRY(CLASET' (fn cs => force_tac |
0be2c98f15a7
replaced best_tac by force_tac, allowing some arithmetic reasoning
paulson
parents:
7262
diff
changeset
|
96 |
(cs addSDs [not0_implies_Suc], ss)) 1), |
6498 | 97 |
simplifier = Rules.simpl_conv ss []}; |
2112 | 98 |
|
99 |
||
100 |
||
6498 | 101 |
val concl = #2 o Rules.dest_thm; |
2112 | 102 |
|
3191 | 103 |
(*--------------------------------------------------------------------------- |
104 |
* Postprocess a definition made by "define". This is a separate stage of |
|
105 |
* processing from the definition stage. |
|
2112 | 106 |
*---------------------------------------------------------------------------*) |
6498 | 107 |
local |
108 |
structure R = Rules |
|
109 |
structure U = Utils |
|
2112 | 110 |
|
6498 | 111 |
(* The rest of these local definitions are for the tricky nested case *) |
112 |
val solved = not o U.can S.dest_eq o #2 o S.strip_forall o concl |
|
2112 | 113 |
|
6498 | 114 |
fun id_thm th = |
115 |
let val {lhs,rhs} = S.dest_eq(#2(S.strip_forall(#2 (R.dest_thm th)))) |
|
116 |
in lhs aconv rhs |
|
117 |
end handle _ => false |
|
2112 | 118 |
|
6498 | 119 |
fun prover s = prove_goal HOL.thy s (fn _ => [fast_tac HOL_cs 1]); |
120 |
val P_imp_P_iff_True = prover "P --> (P= True)" RS mp; |
|
121 |
val P_imp_P_eq_True = P_imp_P_iff_True RS eq_reflection; |
|
122 |
fun mk_meta_eq r = case concl_of r of |
|
123 |
Const("==",_)$_$_ => r |
|
124 |
| _ $(Const("op =",_)$_$_) => r RS eq_reflection |
|
125 |
| _ => r RS P_imp_P_eq_True |
|
3405 | 126 |
|
6498 | 127 |
(*Is this the best way to invoke the simplifier??*) |
128 |
fun rewrite L = rewrite_rule (map mk_meta_eq (filter(not o id_thm) L)) |
|
2112 | 129 |
|
6498 | 130 |
fun join_assums th = |
131 |
let val {sign,...} = rep_thm th |
|
132 |
val tych = cterm_of sign |
|
133 |
val {lhs,rhs} = S.dest_eq(#2 (S.strip_forall (concl th))) |
|
134 |
val cntxtl = (#1 o S.strip_imp) lhs (* cntxtl should = cntxtr *) |
|
135 |
val cntxtr = (#1 o S.strip_imp) rhs (* but union is solider *) |
|
136 |
val cntxt = gen_union (op aconv) (cntxtl, cntxtr) |
|
137 |
in |
|
138 |
R.GEN_ALL |
|
139 |
(R.DISCH_ALL |
|
140 |
(rewrite (map (R.ASSUME o tych) cntxt) (R.SPEC_ALL th))) |
|
141 |
end |
|
142 |
val gen_all = S.gen_all |
|
143 |
in |
|
144 |
fun proof_stage ss theory {f, R, rules, full_pats_TCs, TCs} = |
|
6524 | 145 |
let val dummy = message "Proving induction theorem ..." |
6498 | 146 |
val ind = Prim.mk_induction theory |
147 |
{fconst=f, R=R, SV=[], pat_TCs_list=full_pats_TCs} |
|
8817 | 148 |
val dummy = message "Postprocessing ..."; |
6498 | 149 |
val {rules, induction, nested_tcs} = |
150 |
std_postprocessor ss theory {rules=rules, induction=ind, TCs=TCs} |
|
151 |
in |
|
152 |
case nested_tcs |
|
8817 | 153 |
of [] => {induction=induction, rules=rules,tcs=[]} |
6524 | 154 |
| L => let val dummy = message "Simplifying nested TCs ..." |
6498 | 155 |
val (solved,simplified,stubborn) = |
156 |
U.itlist (fn th => fn (So,Si,St) => |
|
157 |
if (id_thm th) then (So, Si, th::St) else |
|
158 |
if (solved th) then (th::So, Si, St) |
|
159 |
else (So, th::Si, St)) nested_tcs ([],[],[]) |
|
160 |
val simplified' = map join_assums simplified |
|
161 |
val rewr = full_simplify (ss addsimps (solved @ simplified')); |
|
162 |
val induction' = rewr induction |
|
163 |
and rules' = rewr rules |
|
164 |
in |
|
165 |
{induction = induction', |
|
166 |
rules = rules', |
|
167 |
tcs = map (gen_all o S.rhs o #2 o S.strip_forall o concl) |
|
168 |
(simplified@stubborn)} |
|
169 |
end |
|
170 |
end; |
|
3191 | 171 |
|
172 |
||
6498 | 173 |
(*lcp: curry the predicate of the induction rule*) |
174 |
fun curry_rule rl = split_rule_var |
|
175 |
(head_of (HOLogic.dest_Trueprop (concl_of rl)), |
|
176 |
rl); |
|
3191 | 177 |
|
6498 | 178 |
(*lcp: put a theorem into Isabelle form, using meta-level connectives*) |
179 |
val meta_outer = |
|
180 |
curry_rule o standard o |
|
181 |
rule_by_tactic (REPEAT |
|
182 |
(FIRSTGOAL (resolve_tac [allI, impI, conjI] |
|
183 |
ORELSE' etac conjE))); |
|
2112 | 184 |
|
6498 | 185 |
(*Strip off the outer !P*) |
186 |
val spec'= read_instantiate [("x","P::?'b=>bool")] spec; |
|
3459
112cbb8301dc
Removal of structure Context and its replacement by a theorem list of
paulson
parents:
3405
diff
changeset
|
187 |
|
6498 | 188 |
fun simplify_defn (ss, tflCongs) (thy,(id,pats)) = |
189 |
let val dummy = deny (id mem (Sign.stamp_names_of (Theory.sign_of thy))) |
|
190 |
("Recursive definition " ^ id ^ |
|
8817 | 191 |
" would clash with the theory of the same name!") (* FIXME !? *) |
6498 | 192 |
val def = freezeT(get_def thy id) RS meta_eq_to_obj_eq |
8622
870a58dd0ddd
the simplification rules returned from TFL are now paired with the row they
nipkow
parents:
8526
diff
changeset
|
193 |
val {theory,rules,rows,TCs,full_pats_TCs} = |
6498 | 194 |
Prim.post_definition (Prim.congs tflCongs) |
195 |
(thy, (def,pats)) |
|
196 |
val {lhs=f,rhs} = S.dest_eq(concl def) |
|
197 |
val (_,[R,_]) = S.strip_comb rhs |
|
198 |
val ss' = ss addsimps Prim.default_simps |
|
199 |
val {induction, rules, tcs} = |
|
200 |
proof_stage ss' theory |
|
201 |
{f = f, R = R, rules = rules, |
|
202 |
full_pats_TCs = full_pats_TCs, |
|
203 |
TCs = TCs} |
|
204 |
val rules' = map (standard o normalize_thm [RSmp]) (R.CONJUNCTS rules) |
|
205 |
in {induct = meta_outer |
|
206 |
(normalize_thm [RSspec,RSmp] (induction RS spec')), |
|
8622
870a58dd0ddd
the simplification rules returned from TFL are now paired with the row they
nipkow
parents:
8526
diff
changeset
|
207 |
rules = ListPair.zip(rules', rows), |
6498 | 208 |
tcs = (termination_goals rules') @ tcs} |
209 |
end |
|
210 |
handle Utils.ERR {mesg,func,module} => |
|
211 |
error (mesg ^ |
|
212 |
"\n (In TFL function " ^ module ^ "." ^ func ^ ")"); |
|
2112 | 213 |
|
3191 | 214 |
(*--------------------------------------------------------------------------- |
7262 | 215 |
* Defining a function with an associated termination relation. |
216 |
*---------------------------------------------------------------------------*) |
|
217 |
fun define_i thy fid R ss_congs eqs = |
|
218 |
let val {functional,pats} = Prim.mk_functional thy eqs |
|
219 |
val thy = Prim.wfrec_definition0 thy (Sign.base_name fid) R functional |
|
220 |
in (thy, simplify_defn ss_congs (thy, (fid, pats))) |
|
221 |
end; |
|
222 |
||
223 |
fun define thy fid R ss_congs seqs = |
|
8817 | 224 |
define_i thy fid (read_term thy R) ss_congs (map (read_term thy) seqs) |
225 |
handle Utils.ERR {mesg,...} => error mesg; |
|
7262 | 226 |
|
227 |
||
228 |
(*--------------------------------------------------------------------------- |
|
3191 | 229 |
* |
6498 | 230 |
* Definitions with synthesized termination relation |
3191 | 231 |
* |
232 |
*---------------------------------------------------------------------------*) |
|
6498 | 233 |
|
234 |
local open USyntax |
|
235 |
in |
|
236 |
fun func_of_cond_eqn tm = |
|
237 |
#1(strip_comb(#lhs(dest_eq(#2 (strip_forall(#2(strip_imp tm))))))) |
|
238 |
end; |
|
239 |
||
6554 | 240 |
fun defer_i thy fid congs eqs = |
241 |
let val {rules,R,theory,full_pats_TCs,SV,...} = |
|
242 |
Prim.lazyR_def thy (Sign.base_name fid) congs eqs |
|
6498 | 243 |
val f = func_of_cond_eqn (concl(R.CONJUNCT1 rules handle _ => rules)) |
8817 | 244 |
val dummy = message "Proving induction theorem ..."; |
6498 | 245 |
val induction = Prim.mk_induction theory |
246 |
{fconst=f, R=R, SV=SV, pat_TCs_list=full_pats_TCs} |
|
247 |
in (theory, |
|
248 |
(*return the conjoined induction rule and recursion equations, |
|
249 |
with assumptions remaining to discharge*) |
|
250 |
standard (induction RS (rules RS conjI))) |
|
251 |
end |
|
252 |
||
6554 | 253 |
fun defer thy fid congs seqs = |
8817 | 254 |
defer_i thy fid congs (map (read_term thy) seqs) |
255 |
handle Utils.ERR {mesg,...} => error mesg; |
|
6498 | 256 |
end; |
257 |
||
2112 | 258 |
end; |