10769
|
1 |
(* Title: TFL/post.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Konrad Slind, Cambridge University Computer Laboratory
|
|
4 |
Copyright 1997 University of Cambridge
|
|
5 |
|
|
6 |
Second part of main module (postprocessing of TFL definitions).
|
|
7 |
*)
|
|
8 |
|
|
9 |
signature TFL =
|
|
10 |
sig
|
|
11 |
val trace: bool ref
|
|
12 |
val quiet_mode: bool ref
|
|
13 |
val message: string -> unit
|
|
14 |
val tgoalw: theory -> thm list -> thm list -> thm list
|
|
15 |
val tgoal: theory -> thm list -> thm list
|
11632
|
16 |
val define_i: bool -> theory -> claset -> simpset -> thm list -> thm list -> xstring ->
|
10769
|
17 |
term -> term list -> theory * {rules: (thm * int) list, induct: thm, tcs: term list}
|
11632
|
18 |
val define: bool -> theory -> claset -> simpset -> thm list -> thm list -> xstring ->
|
10769
|
19 |
string -> string list -> theory * {rules: (thm * int) list, induct: thm, tcs: term list}
|
|
20 |
val defer_i: theory -> thm list -> xstring -> term list -> theory * thm
|
|
21 |
val defer: theory -> thm list -> xstring -> string list -> theory * thm
|
|
22 |
end;
|
|
23 |
|
|
24 |
structure Tfl: TFL =
|
|
25 |
struct
|
|
26 |
|
|
27 |
structure S = USyntax
|
|
28 |
|
|
29 |
|
|
30 |
(* messages *)
|
|
31 |
|
|
32 |
val trace = Prim.trace
|
|
33 |
|
|
34 |
val quiet_mode = ref false;
|
|
35 |
fun message s = if ! quiet_mode then () else writeln s;
|
|
36 |
|
|
37 |
|
|
38 |
(* misc *)
|
|
39 |
|
12341
|
40 |
val read_term = Thm.term_of oo (HOLogic.read_cterm o Theory.sign_of);
|
10769
|
41 |
|
|
42 |
|
|
43 |
(*---------------------------------------------------------------------------
|
|
44 |
* Extract termination goals so that they can be put it into a goalstack, or
|
|
45 |
* have a tactic directly applied to them.
|
|
46 |
*--------------------------------------------------------------------------*)
|
|
47 |
fun termination_goals rules =
|
|
48 |
map (#1 o Type.freeze_thaw o HOLogic.dest_Trueprop)
|
|
49 |
(foldr (fn (th,A) => union_term (prems_of th, A)) (rules, []));
|
|
50 |
|
|
51 |
(*---------------------------------------------------------------------------
|
|
52 |
* Finds the termination conditions in (highly massaged) definition and
|
|
53 |
* puts them into a goalstack.
|
|
54 |
*--------------------------------------------------------------------------*)
|
|
55 |
fun tgoalw thy defs rules =
|
|
56 |
case termination_goals rules of
|
|
57 |
[] => error "tgoalw: no termination conditions to prove"
|
|
58 |
| L => goalw_cterm defs
|
|
59 |
(Thm.cterm_of (Theory.sign_of thy)
|
|
60 |
(HOLogic.mk_Trueprop(USyntax.list_mk_conj L)));
|
|
61 |
|
|
62 |
fun tgoal thy = tgoalw thy [];
|
|
63 |
|
|
64 |
(*---------------------------------------------------------------------------
|
|
65 |
* Three postprocessors are applied to the definition. It
|
|
66 |
* attempts to prove wellfoundedness of the given relation, simplifies the
|
|
67 |
* non-proved termination conditions, and finally attempts to prove the
|
|
68 |
* simplified termination conditions.
|
|
69 |
*--------------------------------------------------------------------------*)
|
11632
|
70 |
fun std_postprocessor strict cs ss wfs =
|
|
71 |
Prim.postprocess strict
|
10769
|
72 |
{wf_tac = REPEAT (ares_tac wfs 1),
|
|
73 |
terminator = asm_simp_tac ss 1
|
12488
|
74 |
THEN TRY (arith_tac 1 ORELSE
|
|
75 |
fast_tac (cs addSDs [not0_implies_Suc] addss ss) 1),
|
10769
|
76 |
simplifier = Rules.simpl_conv ss []};
|
|
77 |
|
|
78 |
|
|
79 |
|
|
80 |
val concl = #2 o Rules.dest_thm;
|
|
81 |
|
|
82 |
(*---------------------------------------------------------------------------
|
|
83 |
* Postprocess a definition made by "define". This is a separate stage of
|
|
84 |
* processing from the definition stage.
|
|
85 |
*---------------------------------------------------------------------------*)
|
|
86 |
local
|
|
87 |
structure R = Rules
|
|
88 |
structure U = Utils
|
|
89 |
|
|
90 |
(* The rest of these local definitions are for the tricky nested case *)
|
|
91 |
val solved = not o can S.dest_eq o #2 o S.strip_forall o concl
|
|
92 |
|
|
93 |
fun id_thm th =
|
|
94 |
let val {lhs,rhs} = S.dest_eq (#2 (S.strip_forall (#2 (R.dest_thm th))));
|
|
95 |
in lhs aconv rhs end
|
|
96 |
handle U.ERR _ => false;
|
|
97 |
|
|
98 |
|
|
99 |
fun prover s = prove_goal HOL.thy s (fn _ => [fast_tac HOL_cs 1]);
|
|
100 |
val P_imp_P_iff_True = prover "P --> (P= True)" RS mp;
|
|
101 |
val P_imp_P_eq_True = P_imp_P_iff_True RS eq_reflection;
|
|
102 |
fun mk_meta_eq r = case concl_of r of
|
|
103 |
Const("==",_)$_$_ => r
|
|
104 |
| _ $(Const("op =",_)$_$_) => r RS eq_reflection
|
|
105 |
| _ => r RS P_imp_P_eq_True
|
|
106 |
|
|
107 |
(*Is this the best way to invoke the simplifier??*)
|
|
108 |
fun rewrite L = rewrite_rule (map mk_meta_eq (filter(not o id_thm) L))
|
|
109 |
|
|
110 |
fun join_assums th =
|
|
111 |
let val {sign,...} = rep_thm th
|
|
112 |
val tych = cterm_of sign
|
|
113 |
val {lhs,rhs} = S.dest_eq(#2 (S.strip_forall (concl th)))
|
|
114 |
val cntxtl = (#1 o S.strip_imp) lhs (* cntxtl should = cntxtr *)
|
|
115 |
val cntxtr = (#1 o S.strip_imp) rhs (* but union is solider *)
|
|
116 |
val cntxt = gen_union (op aconv) (cntxtl, cntxtr)
|
|
117 |
in
|
|
118 |
R.GEN_ALL
|
|
119 |
(R.DISCH_ALL
|
|
120 |
(rewrite (map (R.ASSUME o tych) cntxt) (R.SPEC_ALL th)))
|
|
121 |
end
|
|
122 |
val gen_all = S.gen_all
|
|
123 |
in
|
11632
|
124 |
fun proof_stage strict cs ss wfs theory {f, R, rules, full_pats_TCs, TCs} =
|
10769
|
125 |
let
|
|
126 |
val _ = message "Proving induction theorem ..."
|
|
127 |
val ind = Prim.mk_induction theory {fconst=f, R=R, SV=[], pat_TCs_list=full_pats_TCs}
|
|
128 |
val _ = message "Postprocessing ...";
|
|
129 |
val {rules, induction, nested_tcs} =
|
11632
|
130 |
std_postprocessor strict cs ss wfs theory {rules=rules, induction=ind, TCs=TCs}
|
10769
|
131 |
in
|
|
132 |
case nested_tcs
|
|
133 |
of [] => {induction=induction, rules=rules,tcs=[]}
|
|
134 |
| L => let val dummy = message "Simplifying nested TCs ..."
|
|
135 |
val (solved,simplified,stubborn) =
|
|
136 |
U.itlist (fn th => fn (So,Si,St) =>
|
|
137 |
if (id_thm th) then (So, Si, th::St) else
|
|
138 |
if (solved th) then (th::So, Si, St)
|
|
139 |
else (So, th::Si, St)) nested_tcs ([],[],[])
|
|
140 |
val simplified' = map join_assums simplified
|
|
141 |
val rewr = full_simplify (ss addsimps (solved @ simplified'));
|
|
142 |
val induction' = rewr induction
|
|
143 |
and rules' = rewr rules
|
|
144 |
in
|
|
145 |
{induction = induction',
|
|
146 |
rules = rules',
|
|
147 |
tcs = map (gen_all o S.rhs o #2 o S.strip_forall o concl)
|
|
148 |
(simplified@stubborn)}
|
|
149 |
end
|
|
150 |
end;
|
|
151 |
|
|
152 |
|
|
153 |
(*lcp: curry the predicate of the induction rule*)
|
11038
|
154 |
fun curry_rule rl =
|
|
155 |
SplitRule.split_rule_var (Term.head_of (HOLogic.dest_Trueprop (concl_of rl)), rl);
|
10769
|
156 |
|
|
157 |
(*lcp: put a theorem into Isabelle form, using meta-level connectives*)
|
|
158 |
val meta_outer =
|
11038
|
159 |
curry_rule o standard o
|
|
160 |
rule_by_tactic (REPEAT (FIRSTGOAL (resolve_tac [allI, impI, conjI] ORELSE' etac conjE)));
|
10769
|
161 |
|
|
162 |
(*Strip off the outer !P*)
|
|
163 |
val spec'= read_instantiate [("x","P::?'b=>bool")] spec;
|
|
164 |
|
11632
|
165 |
fun simplify_defn strict thy cs ss congs wfs id pats def0 =
|
10769
|
166 |
let val def = freezeT def0 RS meta_eq_to_obj_eq
|
|
167 |
val {theory,rules,rows,TCs,full_pats_TCs} = Prim.post_definition congs (thy, (def,pats))
|
|
168 |
val {lhs=f,rhs} = S.dest_eq (concl def)
|
|
169 |
val (_,[R,_]) = S.strip_comb rhs
|
|
170 |
val {induction, rules, tcs} =
|
11632
|
171 |
proof_stage strict cs ss wfs theory
|
10769
|
172 |
{f = f, R = R, rules = rules,
|
|
173 |
full_pats_TCs = full_pats_TCs,
|
|
174 |
TCs = TCs}
|
11771
|
175 |
val rules' = map (standard o ObjectLogic.rulify_no_asm) (R.CONJUNCTS rules)
|
|
176 |
in {induct = meta_outer (ObjectLogic.rulify_no_asm (induction RS spec')),
|
10769
|
177 |
rules = ListPair.zip(rules', rows),
|
|
178 |
tcs = (termination_goals rules') @ tcs}
|
|
179 |
end
|
|
180 |
handle U.ERR {mesg,func,module} =>
|
|
181 |
error (mesg ^
|
|
182 |
"\n (In TFL function " ^ module ^ "." ^ func ^ ")");
|
|
183 |
|
|
184 |
(*---------------------------------------------------------------------------
|
|
185 |
* Defining a function with an associated termination relation.
|
|
186 |
*---------------------------------------------------------------------------*)
|
11632
|
187 |
fun define_i strict thy cs ss congs wfs fid R eqs =
|
10769
|
188 |
let val {functional,pats} = Prim.mk_functional thy eqs
|
|
189 |
val (thy, def) = Prim.wfrec_definition0 thy (Sign.base_name fid) R functional
|
11632
|
190 |
in (thy, simplify_defn strict thy cs ss congs wfs fid pats def) end;
|
10769
|
191 |
|
11632
|
192 |
fun define strict thy cs ss congs wfs fid R seqs =
|
|
193 |
define_i strict thy cs ss congs wfs fid (read_term thy R) (map (read_term thy) seqs)
|
10769
|
194 |
handle U.ERR {mesg,...} => error mesg;
|
|
195 |
|
|
196 |
|
|
197 |
(*---------------------------------------------------------------------------
|
|
198 |
*
|
|
199 |
* Definitions with synthesized termination relation
|
|
200 |
*
|
|
201 |
*---------------------------------------------------------------------------*)
|
|
202 |
|
|
203 |
fun func_of_cond_eqn tm =
|
|
204 |
#1 (S.strip_comb (#lhs (S.dest_eq (#2 (S.strip_forall (#2 (S.strip_imp tm)))))));
|
|
205 |
|
|
206 |
fun defer_i thy congs fid eqs =
|
|
207 |
let val {rules,R,theory,full_pats_TCs,SV,...} =
|
|
208 |
Prim.lazyR_def thy (Sign.base_name fid) congs eqs
|
|
209 |
val f = func_of_cond_eqn (concl (R.CONJUNCT1 rules handle U.ERR _ => rules));
|
|
210 |
val dummy = message "Proving induction theorem ...";
|
|
211 |
val induction = Prim.mk_induction theory
|
|
212 |
{fconst=f, R=R, SV=SV, pat_TCs_list=full_pats_TCs}
|
|
213 |
in (theory,
|
|
214 |
(*return the conjoined induction rule and recursion equations,
|
|
215 |
with assumptions remaining to discharge*)
|
|
216 |
standard (induction RS (rules RS conjI)))
|
|
217 |
end
|
|
218 |
|
|
219 |
fun defer thy congs fid seqs =
|
|
220 |
defer_i thy congs fid (map (read_term thy) seqs)
|
|
221 |
handle U.ERR {mesg,...} => error mesg;
|
|
222 |
end;
|
|
223 |
|
|
224 |
end;
|