author | paulson |
Tue, 20 May 1997 11:49:57 +0200 | |
changeset 3245 | 241838c01caf |
parent 3208 | 8336393de482 |
child 3271 | b873969b05d3 |
permissions | -rw-r--r-- |
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
1 |
(*------------------------------------------------------------------------- |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
2 |
there are 3 postprocessors that get applied to the definition: |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
3 |
|
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
4 |
- a wellfoundedness prover (WF_TAC) |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
5 |
- a simplifier (tries to eliminate the language of termination expressions) |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
6 |
- a termination prover |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
7 |
*-------------------------------------------------------------------------*) |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
8 |
|
3191 | 9 |
signature TFL = |
10 |
sig |
|
2112 | 11 |
structure Prim : TFL_sig |
12 |
||
3191 | 13 |
val tgoalw : theory -> thm list -> thm list -> thm list |
14 |
val tgoal: theory -> thm list -> thm list |
|
2112 | 15 |
|
16 |
val WF_TAC : thm list -> tactic |
|
17 |
||
18 |
val simplifier : thm -> thm |
|
19 |
val std_postprocessor : theory |
|
20 |
-> {induction:thm, rules:thm, TCs:term list list} |
|
21 |
-> {induction:thm, rules:thm, nested_tcs:thm list} |
|
22 |
||
3191 | 23 |
val define_i : theory -> term -> term -> theory * (thm * Prim.pattern list) |
24 |
val define : theory -> string -> string list -> theory * Prim.pattern list |
|
2112 | 25 |
|
3191 | 26 |
val simplify_defn : theory * (string * Prim.pattern list) |
27 |
-> {rules:thm list, induct:thm, tcs:term list} |
|
2112 | 28 |
|
3191 | 29 |
(*------------------------------------------------------------------------- |
30 |
val function : theory -> term -> {theory:theory, eq_ind : thm} |
|
31 |
val lazyR_def: theory -> term -> {theory:theory, eqns : thm} |
|
32 |
*-------------------------------------------------------------------------*) |
|
2112 | 33 |
|
34 |
val tflcongs : theory -> thm list |
|
35 |
||
3191 | 36 |
end; |
37 |
||
38 |
||
39 |
structure Tfl: TFL = |
|
2112 | 40 |
struct |
41 |
structure Prim = Prim |
|
3191 | 42 |
structure S = Prim.USyntax |
2112 | 43 |
|
3191 | 44 |
(*--------------------------------------------------------------------------- |
45 |
* Extract termination goals so that they can be put it into a goalstack, or |
|
46 |
* have a tactic directly applied to them. |
|
47 |
*--------------------------------------------------------------------------*) |
|
48 |
fun termination_goals rules = |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
49 |
map (Logic.freeze_vars o HOLogic.dest_Trueprop) |
3191 | 50 |
(foldr (fn (th,A) => union_term (prems_of th, A)) (rules, [])); |
51 |
||
52 |
(*--------------------------------------------------------------------------- |
|
53 |
* Finds the termination conditions in (highly massaged) definition and |
|
54 |
* puts them into a goalstack. |
|
55 |
*--------------------------------------------------------------------------*) |
|
56 |
fun tgoalw thy defs rules = |
|
57 |
let val L = termination_goals rules |
|
2112 | 58 |
open USyntax |
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
59 |
val g = cterm_of (sign_of thy) (HOLogic.mk_Trueprop(list_mk_conj L)) |
2112 | 60 |
in goalw_cterm defs g |
61 |
end; |
|
62 |
||
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
63 |
fun tgoal thy = tgoalw thy []; |
2112 | 64 |
|
3191 | 65 |
(*--------------------------------------------------------------------------- |
66 |
* Simple wellfoundedness prover. |
|
67 |
*--------------------------------------------------------------------------*) |
|
2112 | 68 |
fun WF_TAC thms = REPEAT(FIRST1(map rtac thms)) |
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
69 |
val WFtac = WF_TAC[wf_measure, wf_inv_image, wf_lex_prod, wf_less_than, |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
70 |
wf_pred_list, wf_trancl]; |
2112 | 71 |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
72 |
val terminator = simp_tac(!simpset addsimps [less_Suc_eq, pred_list_def]) 1 |
3208 | 73 |
THEN TRY(best_tac |
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
74 |
(!claset addSDs [not0_implies_Suc] |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
75 |
addss (!simpset)) 1); |
3191 | 76 |
|
2112 | 77 |
val simpls = [less_eq RS eq_reflection, |
3208 | 78 |
lex_prod_def, rprod_def, measure_def, inv_image_def]; |
2112 | 79 |
|
3191 | 80 |
(*--------------------------------------------------------------------------- |
81 |
* Does some standard things with the termination conditions of a definition: |
|
82 |
* attempts to prove wellfoundedness of the given relation; simplifies the |
|
83 |
* non-proven termination conditions; and finally attempts to prove the |
|
84 |
* simplified termination conditions. |
|
85 |
*--------------------------------------------------------------------------*) |
|
2112 | 86 |
val std_postprocessor = Prim.postprocess{WFtac = WFtac, |
87 |
terminator = terminator, |
|
88 |
simplifier = Prim.Rules.simpl_conv simpls}; |
|
89 |
||
3208 | 90 |
val simplifier = rewrite_rule (simpls @ #simps(rep_ss (!simpset)) @ |
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
91 |
[pred_list_def]); |
3191 | 92 |
|
2112 | 93 |
fun tflcongs thy = Prim.Context.read() @ (#case_congs(Thry.extract_info thy)); |
94 |
||
95 |
||
96 |
val concl = #2 o Prim.Rules.dest_thm; |
|
97 |
||
98 |
(*--------------------------------------------------------------------------- |
|
3191 | 99 |
* Defining a function with an associated termination relation. |
100 |
*---------------------------------------------------------------------------*) |
|
101 |
fun define_i thy R eqs = |
|
102 |
let val dummy = require_thy thy "WF_Rel" "recursive function definitions"; |
|
103 |
||
104 |
val {functional,pats} = Prim.mk_functional thy eqs |
|
105 |
val (thm,thry) = Prim.wfrec_definition0 thy R functional |
|
106 |
in (thry,(thm,pats)) |
|
107 |
end; |
|
108 |
||
109 |
(*lcp's version: takes strings; doesn't return "thm" |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
110 |
(whose signature is a draft and therefore useless) *) |
3191 | 111 |
fun define thy R eqs = |
112 |
let fun read thy = readtm (sign_of thy) (TVar(("DUMMY",0),[])) |
|
113 |
val (thy',(_,pats)) = |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
114 |
define_i thy (read thy R) |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
115 |
(fold_bal (app Ind_Syntax.conj) (map (read thy) eqs)) |
3191 | 116 |
in (thy',pats) end |
117 |
handle Utils.ERR {mesg,...} => error mesg; |
|
118 |
||
119 |
(*--------------------------------------------------------------------------- |
|
120 |
* Postprocess a definition made by "define". This is a separate stage of |
|
121 |
* processing from the definition stage. |
|
2112 | 122 |
*---------------------------------------------------------------------------*) |
123 |
local |
|
124 |
structure R = Prim.Rules |
|
125 |
structure U = Utils |
|
126 |
||
3191 | 127 |
(* The rest of these local definitions are for the tricky nested case *) |
2112 | 128 |
val solved = not o U.can S.dest_eq o #2 o S.strip_forall o concl |
129 |
||
130 |
fun id_thm th = |
|
131 |
let val {lhs,rhs} = S.dest_eq(#2(S.strip_forall(#2 (R.dest_thm th)))) |
|
132 |
in S.aconv lhs rhs |
|
133 |
end handle _ => false |
|
134 |
||
135 |
fun prover s = prove_goal HOL.thy s (fn _ => [fast_tac HOL_cs 1]); |
|
136 |
val P_imp_P_iff_True = prover "P --> (P= True)" RS mp; |
|
137 |
val P_imp_P_eq_True = P_imp_P_iff_True RS eq_reflection; |
|
138 |
fun mk_meta_eq r = case concl_of r of |
|
139 |
Const("==",_)$_$_ => r |
|
140 |
| _$(Const("op =",_)$_$_) => r RS eq_reflection |
|
141 |
| _ => r RS P_imp_P_eq_True |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
142 |
fun rewrite L = rewrite_rule (map mk_meta_eq (filter(not o id_thm) L)) |
3191 | 143 |
fun reducer thl = rewrite (map standard thl @ #simps(rep_ss HOL_ss)) |
2112 | 144 |
|
145 |
fun join_assums th = |
|
146 |
let val {sign,...} = rep_thm th |
|
147 |
val tych = cterm_of sign |
|
148 |
val {lhs,rhs} = S.dest_eq(#2 (S.strip_forall (concl th))) |
|
149 |
val cntxtl = (#1 o S.strip_imp) lhs (* cntxtl should = cntxtr *) |
|
150 |
val cntxtr = (#1 o S.strip_imp) rhs (* but union is solider *) |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
151 |
val cntxt = gen_union (op aconv) (cntxtl, cntxtr) |
2112 | 152 |
in |
3191 | 153 |
R.GEN_ALL |
154 |
(R.DISCH_ALL |
|
155 |
(rewrite (map (R.ASSUME o tych) cntxt) (R.SPEC_ALL th))) |
|
2112 | 156 |
end |
157 |
val gen_all = S.gen_all |
|
158 |
in |
|
3191 | 159 |
(*--------------------------------------------------------------------------- |
160 |
* The "reducer" argument is |
|
161 |
* (fn thl => rewrite (map standard thl @ #simps(rep_ss HOL_ss))); |
|
162 |
*---------------------------------------------------------------------------*) |
|
163 |
fun proof_stage theory reducer {f, R, rules, full_pats_TCs, TCs} = |
|
164 |
let val dummy = output(std_out, "Proving induction theorem.. ") |
|
165 |
val ind = Prim.mk_induction theory f R full_pats_TCs |
|
166 |
val dummy = output(std_out, "Proved induction theorem.\n") |
|
167 |
val pp = std_postprocessor theory |
|
168 |
val dummy = output(std_out, "Postprocessing.. ") |
|
169 |
val {rules,induction,nested_tcs} = pp{rules=rules,induction=ind,TCs=TCs} |
|
170 |
in |
|
171 |
case nested_tcs |
|
172 |
of [] => (output(std_out, "Postprocessing done.\n"); |
|
173 |
{induction=induction, rules=rules,tcs=[]}) |
|
174 |
| L => let val dummy = output(std_out, "Simplifying nested TCs.. ") |
|
2112 | 175 |
val (solved,simplified,stubborn) = |
176 |
U.itlist (fn th => fn (So,Si,St) => |
|
177 |
if (id_thm th) then (So, Si, th::St) else |
|
178 |
if (solved th) then (th::So, Si, St) |
|
179 |
else (So, th::Si, St)) nested_tcs ([],[],[]) |
|
180 |
val simplified' = map join_assums simplified |
|
181 |
val induction' = reducer (solved@simplified') induction |
|
182 |
val rules' = reducer (solved@simplified') rules |
|
3191 | 183 |
val dummy = output(std_out, "Postprocessing done.\n") |
2112 | 184 |
in |
185 |
{induction = induction', |
|
186 |
rules = rules', |
|
3191 | 187 |
tcs = map (gen_all o S.rhs o #2 o S.strip_forall o concl) |
188 |
(simplified@stubborn)} |
|
2112 | 189 |
end |
3191 | 190 |
end handle (e as Utils.ERR _) => Utils.Raise e |
191 |
| e => print_exn e; |
|
192 |
||
193 |
||
194 |
(*lcp: put a theorem into Isabelle form, using meta-level connectives*) |
|
195 |
val meta_outer = |
|
196 |
standard o rule_by_tactic (REPEAT_FIRST (resolve_tac [allI, impI, conjI])); |
|
197 |
||
198 |
||
199 |
(*Strip off the outer !P*) |
|
200 |
val spec'= read_instantiate [("x","P::?'b=>bool")] spec; |
|
2112 | 201 |
|
202 |
||
3191 | 203 |
fun simplify_defn (thy,(id,pats)) = |
3208 | 204 |
let val dummy = deny (id mem map ! (stamps_of_thy thy)) |
205 |
("Recursive definition " ^ id ^ |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
206 |
" would clash with the theory of the same name!") |
3208 | 207 |
val def = freezeT(get_def thy id RS meta_eq_to_obj_eq) |
3191 | 208 |
val {theory,rules,TCs,full_pats_TCs,patterns} = |
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
209 |
Prim.post_definition (thy,(def,pats)) |
3191 | 210 |
val {lhs=f,rhs} = S.dest_eq(concl def) |
211 |
val (_,[R,_]) = S.strip_comb rhs |
|
212 |
val {induction, rules, tcs} = |
|
213 |
proof_stage theory reducer |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
214 |
{f = f, R = R, rules = rules, |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
215 |
full_pats_TCs = full_pats_TCs, |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
216 |
TCs = TCs} |
3191 | 217 |
val rules' = map (standard o normalize_thm [RSmp]) (R.CONJUNCTS rules) |
218 |
in {induct = meta_outer |
|
3245
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
219 |
(normalize_thm [RSspec,RSmp] (induction RS spec')), |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
220 |
rules = rules', |
241838c01caf
Removal of redundant code (unused or already present in Isabelle.
paulson
parents:
3208
diff
changeset
|
221 |
tcs = (termination_goals rules') @ tcs} |
3191 | 222 |
end |
223 |
handle Utils.ERR {mesg,...} => error mesg |
|
2112 | 224 |
end; |
225 |
||
3191 | 226 |
(*--------------------------------------------------------------------------- |
227 |
* |
|
228 |
* Definitions with synthesized termination relation temporarily |
|
229 |
* deleted -- it's not clear how to integrate this facility with |
|
230 |
* the Isabelle theory file scheme, which restricts |
|
231 |
* inference at theory-construction time. |
|
232 |
* |
|
2112 | 233 |
|
3208 | 234 |
local structure R = Prim.Rules |
2112 | 235 |
in |
236 |
fun function theory eqs = |
|
3208 | 237 |
let val dummy = prs "Making definition.. " |
2112 | 238 |
val {rules,R,theory,full_pats_TCs,...} = Prim.lazyR_def theory eqs |
239 |
val f = func_of_cond_eqn (concl(R.CONJUNCT1 rules handle _ => rules)) |
|
3208 | 240 |
val dummy = prs "Definition made.\n" |
241 |
val dummy = prs "Proving induction theorem.. " |
|
2112 | 242 |
val induction = Prim.mk_induction theory f R full_pats_TCs |
3208 | 243 |
val dummy = prs "Induction theorem proved.\n" |
2112 | 244 |
in {theory = theory, |
245 |
eq_ind = standard (induction RS (rules RS conjI))} |
|
246 |
end |
|
247 |
handle (e as Utils.ERR _) => Utils.Raise e |
|
248 |
| e => print_exn e |
|
249 |
end; |
|
250 |
||
251 |
||
252 |
fun lazyR_def theory eqs = |
|
253 |
let val {rules,theory, ...} = Prim.lazyR_def theory eqs |
|
254 |
in {eqns=rules, theory=theory} |
|
255 |
end |
|
256 |
handle (e as Utils.ERR _) => Utils.Raise e |
|
257 |
| e => print_exn e; |
|
3191 | 258 |
* |
259 |
* |
|
260 |
*---------------------------------------------------------------------------*) |
|
261 |
||
262 |
||
2112 | 263 |
|
264 |
||
3191 | 265 |
(*--------------------------------------------------------------------------- |
266 |
* Install the basic context notions. Others (for nat and list and prod) |
|
267 |
* have already been added in thry.sml |
|
268 |
*---------------------------------------------------------------------------*) |
|
269 |
val () = Prim.Context.write[Thms.LET_CONG, Thms.COND_CONG]; |
|
2112 | 270 |
|
271 |
end; |