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 close_result: thm -> 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: Proof.context -> string list -> term list -> term ->
28 ({prems: thm list, context: Proof.context} -> tactic) -> thm
29 val prove_global: theory -> string list -> term list -> term -> (thm list -> tactic) -> thm
30 val extract: int -> int -> thm -> thm Seq.seq
31 val retrofit: int -> int -> thm -> thm -> thm Seq.seq
32 val conjunction_tac: int -> tactic
33 val precise_conjunction_tac: int -> int -> tactic
34 val recover_conjunction_tac: tactic
35 val asm_rewrite_goal_tac: bool * bool * bool -> (simpset -> tactic) -> simpset -> int -> tactic
36 val rewrite_goal_tac: thm list -> int -> tactic
37 val norm_hhf_tac: int -> tactic
38 val compose_hhf: thm -> int -> thm -> thm Seq.seq
39 val compose_hhf_tac: thm -> int -> tactic
40 val comp_hhf: thm -> thm -> thm
41 val assume_rule_tac: Proof.context -> int -> tactic
44 structure Goal: GOAL =
54 let val A = #1 (Thm.dest_implies (Thm.cprop_of Drule.protectI))
55 in fn C => Thm.instantiate ([], [(A, C)]) Drule.protectI end;
62 fun protect th = th COMP_INCR Drule.protectI;
66 ---------------- (conclude)
70 (case SINGLE (Thm.compose_no_flatten false (th, Thm.nprems_of th) 1)
71 (Drule.incr_indexes th Drule.protectD) of
73 | NONE => raise THM ("Failed to conclude goal", 0, [th]));
81 (case Thm.nprems_of th of
83 | n => raise THM ("Proof failed.\n" ^
84 Pretty.string_of (Pretty.chunks (Display.pretty_goals n th)) ^
85 ("\n" ^ string_of_int n ^ " unsolved goal(s)!"), 0, [th]));
95 #> MetaSimplifier.norm_hhf_protect
97 #> Drule.zero_var_indexes;
101 #> Drule.close_derivation;
105 (** tactical theorem proving **)
107 (* prove_internal -- minimal checks, no normalization of result! *)
109 fun prove_internal casms cprop tac =
110 (case SINGLE (tac (map Assumption.assume casms)) (init cprop) of
111 SOME th => Drule.implies_intr_list casms (finish th)
112 | NONE => error "Tactic failed.");
117 fun prove_multi ctxt xs asms props tac =
119 val thy = ProofContext.theory_of ctxt;
120 val string_of_term = Sign.string_of_term thy;
122 fun err msg = cat_error msg
123 ("The error(s) above occurred for the goal statement:\n" ^
124 string_of_term (Logic.list_implies (asms, Logic.mk_conjunction_list props)));
126 fun cert_safe t = Thm.cterm_of thy (Envir.beta_norm (Term.no_dummy_patterns t))
127 handle TERM (msg, _) => err msg | TYPE (msg, _, _) => err msg;
128 val casms = map cert_safe asms;
129 val cprops = map cert_safe props;
131 val (prems, ctxt') = ctxt
132 |> Variable.add_fixes_direct xs
133 |> fold Variable.declare_internal (asms @ props)
134 |> Assumption.add_assumes casms;
136 val goal = init (Conjunction.mk_conjunction_balanced cprops);
138 (case SINGLE (tac {prems = prems, context = ctxt'}) goal of
139 NONE => err "Tactic failed."
141 val results = Conjunction.elim_balanced (length props) (finish res)
142 handle THM (msg, _, _) => err msg;
143 val _ = Unify.matches_list thy (map Thm.term_of cprops) (map Thm.prop_of results)
144 orelse err ("Proved a different theorem: " ^ string_of_term (Thm.prop_of res));
147 |> map (Assumption.export false ctxt' ctxt)
148 |> Variable.export ctxt' ctxt
149 |> map Drule.zero_var_indexes
155 fun prove ctxt xs asms prop tac = hd (prove_multi ctxt xs asms [prop] tac);
157 fun prove_global thy xs asms prop tac =
158 Drule.standard (prove (ProofContext.init thy) xs asms prop (fn {prems, ...} => tac prems));
162 (** goal structure **)
167 (if i < 1 orelse n < 1 orelse i + n - 1 > Thm.nprems_of st then Seq.empty
168 else if n = 1 then Seq.single (Thm.cprem_of st i)
170 Seq.single (Conjunction.mk_conjunction_balanced (map (Thm.cprem_of st) (i upto i + n - 1))))
171 |> Seq.map (Thm.adjust_maxidx_cterm ~1 #> init);
173 fun retrofit i n st' st =
175 else st |> Drule.with_subgoal i (Conjunction.uncurry_balanced n))
176 |> Thm.compose_no_flatten false (conclude st', Thm.nprems_of st') i;
178 fun SELECT_GOAL tac i st =
179 if Thm.nprems_of st = 1 andalso i = 1 then tac st
180 else Seq.lifts (retrofit i 1) (Seq.maps tac (extract i 1 st)) st;
185 fun precise_conjunction_tac 0 i = eq_assume_tac i
186 | precise_conjunction_tac 1 i = SUBGOAL (K all_tac) i
187 | precise_conjunction_tac n i = PRIMITIVE (Drule.with_subgoal i (Conjunction.curry_balanced n));
189 val adhoc_conjunction_tac = REPEAT_ALL_NEW
190 (SUBGOAL (fn (goal, i) =>
191 if can Logic.dest_conjunction goal then rtac Conjunction.conjunctionI i
194 val conjunction_tac = SUBGOAL (fn (goal, i) =>
195 precise_conjunction_tac (length (Logic.dest_conjunctions goal)) i ORELSE
196 TRY (adhoc_conjunction_tac i));
198 val recover_conjunction_tac = PRIMITIVE (fn th =>
199 Conjunction.uncurry_balanced (Thm.nprems_of th) th);
201 fun PRECISE_CONJUNCTS n tac =
202 SELECT_GOAL (precise_conjunction_tac n 1
204 THEN recover_conjunction_tac);
207 SELECT_GOAL (conjunction_tac 1
209 THEN recover_conjunction_tac);
214 (*Rewrite subgoal i only. SELECT_GOAL avoids inefficiencies in goals_conv.*)
215 fun asm_rewrite_goal_tac mode prover_tac ss =
217 (PRIMITIVE (MetaSimplifier.rewrite_goal_rule mode (SINGLE o prover_tac) ss 1));
219 fun rewrite_goal_tac rews =
220 let val ss = MetaSimplifier.empty_ss addsimps rews in
221 fn i => fn st => asm_rewrite_goal_tac (true, false, false) (K no_tac)
222 (MetaSimplifier.theory_context (Thm.theory_of_thm st) ss) i st
226 (* hhf normal form *)
229 rtac Drule.asm_rl (*cheap approximation -- thanks to builtin Logic.flatten_params*)
230 THEN' SUBGOAL (fn (t, i) =>
231 if Drule.is_norm_hhf t then all_tac
232 else rewrite_goal_tac [Drule.norm_hhf_eq] i);
234 fun compose_hhf tha i thb =
235 Thm.bicompose false (false, Drule.lift_all (Thm.cprem_of thb i) tha, 0) i thb;
237 fun compose_hhf_tac th i = PRIMSEQ (compose_hhf th i);
239 fun comp_hhf tha thb =
240 (case Seq.chop 2 (compose_hhf tha 1 thb) of
242 | ([], _) => raise THM ("comp_hhf: no unifiers", 1, [tha, thb])
243 | _ => raise THM ("comp_hhf: multiple unifiers", 1, [tha, thb]));
246 (* non-atomic goal assumptions *)
248 fun non_atomic (Const ("==>", _) $ _ $ _) = true
249 | non_atomic (Const ("all", _) $ _) = true
250 | non_atomic _ = false;
252 fun assume_rule_tac ctxt = norm_hhf_tac THEN' CSUBGOAL (fn (goal, i) =>
254 val ((_, goal'), ctxt') = Variable.focus goal ctxt;
255 val goal'' = Drule.cterm_rule (singleton (Variable.export ctxt' ctxt)) goal';
256 val Rs = filter (non_atomic o Thm.term_of) (Drule.strip_imp_prems goal'');
257 val tacs = Rs |> map (fn R =>
258 Tactic.etac (MetaSimplifier.norm_hhf (Thm.trivial R)) THEN_ALL_NEW assume_tac);
259 in fold_rev (curry op APPEND') tacs (K no_tac) i end);
263 structure BasicGoal: BASIC_GOAL = Goal;