3 Author: Makarius and Lawrence C Paulson
5 Goals in tactical theorem proving.
10 val SELECT_GOAL: tactic -> int -> tactic
11 val CONJUNCTS: tactic -> int -> tactic
12 val PRECISE_CONJUNCTS: int -> tactic -> int -> tactic
18 val init: cterm -> thm
19 val protect: thm -> thm
20 val conclude: thm -> thm
21 val finish: thm -> thm
22 val norm_result: thm -> thm
23 val future_result: Proof.context -> thm future -> term -> thm
24 val prove_internal: cterm list -> cterm -> (thm list -> tactic) -> thm
25 val prove_multi: Proof.context -> string list -> term list -> term list ->
26 ({prems: thm list, context: Proof.context} -> tactic) -> thm list
27 val prove_future: Proof.context -> string list -> term list -> term ->
28 ({prems: thm list, context: Proof.context} -> tactic) -> thm
29 val prove: Proof.context -> string list -> term list -> term ->
30 ({prems: thm list, context: Proof.context} -> tactic) -> thm
31 val prove_global: theory -> string list -> term list -> term ->
32 ({prems: thm list, context: Proof.context} -> tactic) -> thm
33 val extract: int -> int -> thm -> thm Seq.seq
34 val retrofit: int -> int -> thm -> thm -> thm Seq.seq
35 val conjunction_tac: int -> tactic
36 val precise_conjunction_tac: int -> int -> tactic
37 val recover_conjunction_tac: tactic
38 val norm_hhf_tac: int -> tactic
39 val compose_hhf_tac: thm -> int -> tactic
40 val assume_rule_tac: Proof.context -> int -> tactic
43 structure Goal: GOAL =
53 let val A = #1 (Thm.dest_implies (Thm.cprop_of Drule.protectI))
54 in fn C => Thm.instantiate ([], [(A, C)]) Drule.protectI end;
61 fun protect th = th COMP_INCR Drule.protectI;
65 ---------------- (conclude)
69 (case SINGLE (Thm.compose_no_flatten false (th, Thm.nprems_of th) 1)
70 (Drule.incr_indexes th Drule.protectD) of
72 | NONE => raise THM ("Failed to conclude goal", 0, [th]));
80 (case Thm.nprems_of th of
82 | n => raise THM ("Proof failed.\n" ^
83 Pretty.string_of (Pretty.chunks (Display.pretty_goals n th)) ^
84 ("\n" ^ string_of_int n ^ " unsolved goal(s)!"), 0, [th]));
94 #> MetaSimplifier.norm_hhf_protect
96 #> Drule.zero_var_indexes;
101 fun future_result ctxt result prop =
103 val thy = ProofContext.theory_of ctxt;
104 val _ = Context.reject_draft thy;
105 val cert = Thm.cterm_of thy;
106 val certT = Thm.ctyp_of thy;
108 val assms = Assumption.assms_of ctxt;
109 val As = map Thm.term_of assms;
111 val xs = map Free (fold Term.add_frees (prop :: As) []);
112 val fixes = map cert xs;
114 val tfrees = fold Term.add_tfrees (prop :: As) [];
115 val instT = map (fn (a, S) => (certT (TVar ((a, 0), S)), certT (TFree (a, S)))) tfrees;
118 Term.map_types Logic.varifyT (fold_rev Logic.all xs (Logic.list_implies (As, prop)));
119 val global_result = result |> Future.map
120 (Thm.adjust_maxidx_thm ~1 #>
121 Drule.implies_intr_list assms #>
122 Drule.forall_intr_list fixes #>
123 Thm.generalize (map #1 tfrees, []) 0);
125 Thm.future global_result (cert global_prop)
126 |> Thm.instantiate (instT, [])
127 |> Drule.forall_elim_list fixes
128 |> fold (Thm.elim_implies o Thm.assume) assms;
133 (** tactical theorem proving **)
135 (* prove_internal -- minimal checks, no normalization of result! *)
137 fun prove_internal casms cprop tac =
138 (case SINGLE (tac (map Assumption.assume casms)) (init cprop) of
139 SOME th => Drule.implies_intr_list casms (finish th)
140 | NONE => error "Tactic failed.");
143 (* prove_common etc. *)
145 fun prove_common immediate ctxt xs asms props tac =
147 val thy = ProofContext.theory_of ctxt;
148 val string_of_term = Syntax.string_of_term ctxt;
150 val pos = Position.thread_data ();
151 fun err msg = cat_error msg
152 ("The error(s) above occurred for the goal statement:\n" ^
153 string_of_term (Logic.list_implies (asms, Logic.mk_conjunction_list props)) ^
154 (case Position.str_of pos of "" => "" | s => "\n" ^ s));
156 fun cert_safe t = Thm.cterm_of thy (Envir.beta_norm (Term.no_dummy_patterns t))
157 handle TERM (msg, _) => err msg | TYPE (msg, _, _) => err msg;
158 val casms = map cert_safe asms;
159 val cprops = map cert_safe props;
161 val (prems, ctxt') = ctxt
162 |> Variable.add_fixes_direct xs
163 |> fold Variable.declare_term (asms @ props)
164 |> Assumption.add_assumes casms
165 ||> Variable.set_body true;
166 val sorts = Variable.sorts_of ctxt';
168 val stmt = Thm.weaken_sorts sorts (Conjunction.mk_conjunction_balanced cprops);
171 (case SINGLE (tac {prems = prems, context = ctxt'}) (init stmt) of
172 NONE => err "Tactic failed."
174 let val res = finish st handle THM (msg, _, _) => err msg in
175 if Unify.matches_list thy [Thm.term_of stmt] [Thm.prop_of res]
176 then Thm.check_shyps sorts res
177 else err ("Proved a different theorem: " ^ string_of_term (Thm.prop_of res))
180 if immediate orelse #maxidx (Thm.rep_cterm stmt) >= 0 orelse not (Future.enabled ())
182 else future_result ctxt' (Future.fork_pri 1 result) (Thm.term_of stmt);
184 Conjunction.elim_balanced (length props) res
185 |> map (Assumption.export false ctxt' ctxt)
186 |> Variable.export ctxt' ctxt
187 |> map Drule.zero_var_indexes
190 val prove_multi = prove_common true;
192 fun prove_future ctxt xs asms prop tac = hd (prove_common false ctxt xs asms [prop] tac);
193 fun prove ctxt xs asms prop tac = hd (prove_common true ctxt xs asms [prop] tac);
195 fun prove_global thy xs asms prop tac =
196 Drule.standard (prove (ProofContext.init thy) xs asms prop tac);
200 (** goal structure **)
205 (if i < 1 orelse n < 1 orelse i + n - 1 > Thm.nprems_of st then Seq.empty
206 else if n = 1 then Seq.single (Thm.cprem_of st i)
208 Seq.single (Conjunction.mk_conjunction_balanced (map (Thm.cprem_of st) (i upto i + n - 1))))
209 |> Seq.map (Thm.adjust_maxidx_cterm ~1 #> init);
211 fun retrofit i n st' st =
213 else st |> Drule.with_subgoal i (Conjunction.uncurry_balanced n))
214 |> Thm.compose_no_flatten false (conclude st', Thm.nprems_of st') i;
216 fun SELECT_GOAL tac i st =
217 if Thm.nprems_of st = 1 andalso i = 1 then tac st
218 else Seq.lifts (retrofit i 1) (Seq.maps tac (extract i 1 st)) st;
223 fun precise_conjunction_tac 0 i = eq_assume_tac i
224 | precise_conjunction_tac 1 i = SUBGOAL (K all_tac) i
225 | precise_conjunction_tac n i = PRIMITIVE (Drule.with_subgoal i (Conjunction.curry_balanced n));
227 val adhoc_conjunction_tac = REPEAT_ALL_NEW
228 (SUBGOAL (fn (goal, i) =>
229 if can Logic.dest_conjunction goal then rtac Conjunction.conjunctionI i
232 val conjunction_tac = SUBGOAL (fn (goal, i) =>
233 precise_conjunction_tac (length (Logic.dest_conjunctions goal)) i ORELSE
234 TRY (adhoc_conjunction_tac i));
236 val recover_conjunction_tac = PRIMITIVE (fn th =>
237 Conjunction.uncurry_balanced (Thm.nprems_of th) th);
239 fun PRECISE_CONJUNCTS n tac =
240 SELECT_GOAL (precise_conjunction_tac n 1
242 THEN recover_conjunction_tac);
245 SELECT_GOAL (conjunction_tac 1
247 THEN recover_conjunction_tac);
250 (* hhf normal form *)
253 rtac Drule.asm_rl (*cheap approximation -- thanks to builtin Logic.flatten_params*)
254 THEN' SUBGOAL (fn (t, i) =>
255 if Drule.is_norm_hhf t then all_tac
256 else MetaSimplifier.rewrite_goal_tac Drule.norm_hhf_eqs i);
258 fun compose_hhf_tac th i st =
259 PRIMSEQ (Thm.bicompose false (false, Drule.lift_all (Thm.cprem_of st i) th, 0) i) st;
262 (* non-atomic goal assumptions *)
264 fun non_atomic (Const ("==>", _) $ _ $ _) = true
265 | non_atomic (Const ("all", _) $ _) = true
266 | non_atomic _ = false;
268 fun assume_rule_tac ctxt = norm_hhf_tac THEN' CSUBGOAL (fn (goal, i) =>
270 val ((_, goal'), ctxt') = Variable.focus goal ctxt;
271 val goal'' = Drule.cterm_rule (singleton (Variable.export ctxt' ctxt)) goal';
272 val Rs = filter (non_atomic o Thm.term_of) (Drule.strip_imp_prems goal'');
273 val tacs = Rs |> map (fn R =>
274 Tactic.etac (MetaSimplifier.norm_hhf (Thm.trivial R)) THEN_ALL_NEW assume_tac);
275 in fold_rev (curry op APPEND') tacs (K no_tac) i end);
279 structure BasicGoal: BASIC_GOAL = Goal;