| author | wenzelm | 
| Sat, 27 Oct 2001 00:07:19 +0200 | |
| changeset 11961 | e5911a25d094 | 
| parent 11929 | a77ad6c86564 | 
| child 11970 | e7eedbd2c8ca | 
| permissions | -rw-r--r-- | 
| 10805 | 1  | 
(* Title: Pure/tactic.ML  | 
| 0 | 2  | 
ID: $Id$  | 
| 10805 | 3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 0 | 4  | 
Copyright 1991 University of Cambridge  | 
5  | 
||
| 10805 | 6  | 
Tactics.  | 
| 0 | 7  | 
*)  | 
8  | 
||
| 11774 | 9  | 
signature BASIC_TACTIC =  | 
10  | 
sig  | 
|
| 10805 | 11  | 
val ares_tac : thm list -> int -> tactic  | 
| 11671 | 12  | 
val asm_rewrite_goal_tac: bool*bool*bool ->  | 
13  | 
(MetaSimplifier.meta_simpset -> tactic) -> MetaSimplifier.meta_simpset -> int -> tactic  | 
|
| 10805 | 14  | 
val assume_tac : int -> tactic  | 
15  | 
val atac : int ->tactic  | 
|
| 10817 | 16  | 
val bimatch_from_nets_tac:  | 
| 1501 | 17  | 
(int*(bool*thm)) Net.net * (int*(bool*thm)) Net.net -> int -> tactic  | 
| 10805 | 18  | 
val bimatch_tac : (bool*thm)list -> int -> tactic  | 
| 10817 | 19  | 
val biresolution_from_nets_tac:  | 
| 10805 | 20  | 
        ('a list -> (bool * thm) list) ->
 | 
21  | 
bool -> 'a Net.net * 'a Net.net -> int -> tactic  | 
|
| 10817 | 22  | 
val biresolve_from_nets_tac:  | 
| 1501 | 23  | 
(int*(bool*thm)) Net.net * (int*(bool*thm)) Net.net -> int -> tactic  | 
| 10805 | 24  | 
val biresolve_tac : (bool*thm)list -> int -> tactic  | 
25  | 
val build_net : thm list -> (int*thm) Net.net  | 
|
| 1501 | 26  | 
val build_netpair: (int*(bool*thm)) Net.net * (int*(bool*thm)) Net.net ->  | 
27  | 
(bool*thm)list -> (int*(bool*thm)) Net.net * (int*(bool*thm)) Net.net  | 
|
| 10817 | 28  | 
val compose_inst_tac : (string*string)list -> (bool*thm*int) ->  | 
| 
3409
 
c0466958df5d
Tidying of signature.  More robust renaming in freeze_thaw.
 
paulson 
parents: 
2814 
diff
changeset
 | 
29  | 
int -> tactic  | 
| 10817 | 30  | 
val compose_tac : (bool * thm * int) -> int -> tactic  | 
| 10805 | 31  | 
val cut_facts_tac : thm list -> int -> tactic  | 
| 10817 | 32  | 
val cut_inst_tac : (string*string)list -> thm -> int -> tactic  | 
| 7491 | 33  | 
val datac : thm -> int -> int -> tactic  | 
| 10805 | 34  | 
val defer_tac : int -> tactic  | 
35  | 
val distinct_subgoals_tac : tactic  | 
|
36  | 
val dmatch_tac : thm list -> int -> tactic  | 
|
37  | 
val dresolve_tac : thm list -> int -> tactic  | 
|
| 10817 | 38  | 
val dres_inst_tac : (string*string)list -> thm -> int -> tactic  | 
| 10805 | 39  | 
val dtac : thm -> int ->tactic  | 
| 7491 | 40  | 
val eatac : thm -> int -> int -> tactic  | 
| 10805 | 41  | 
val etac : thm -> int ->tactic  | 
| 10817 | 42  | 
val eq_assume_tac : int -> tactic  | 
| 10805 | 43  | 
val ematch_tac : thm list -> int -> tactic  | 
44  | 
val eresolve_tac : thm list -> int -> tactic  | 
|
45  | 
val eres_inst_tac : (string*string)list -> thm -> int -> tactic  | 
|
| 7491 | 46  | 
val fatac : thm -> int -> int -> tactic  | 
| 10817 | 47  | 
val filter_prems_tac : (term -> bool) -> int -> tactic  | 
| 10805 | 48  | 
val filter_thms : (term*term->bool) -> int*term*thm list -> thm list  | 
49  | 
val filt_resolve_tac : thm list -> int -> int -> tactic  | 
|
50  | 
val flexflex_tac : tactic  | 
|
51  | 
val fold_goals_tac : thm list -> tactic  | 
|
52  | 
val fold_rule : thm list -> thm -> thm  | 
|
53  | 
val fold_tac : thm list -> tactic  | 
|
| 10817 | 54  | 
val forward_tac : thm list -> int -> tactic  | 
| 10805 | 55  | 
val forw_inst_tac : (string*string)list -> thm -> int -> tactic  | 
56  | 
val ftac : thm -> int ->tactic  | 
|
| 10817 | 57  | 
  val insert_tagged_brl : ('a*(bool*thm)) *
 | 
| 
3409
 
c0466958df5d
Tidying of signature.  More robust renaming in freeze_thaw.
 
paulson 
parents: 
2814 
diff
changeset
 | 
58  | 
                          (('a*(bool*thm))Net.net * ('a*(bool*thm))Net.net) ->
 | 
| 
 
c0466958df5d
Tidying of signature.  More robust renaming in freeze_thaw.
 
paulson 
parents: 
2814 
diff
changeset
 | 
59  | 
                          ('a*(bool*thm))Net.net * ('a*(bool*thm))Net.net
 | 
| 10817 | 60  | 
val delete_tagged_brl : (bool*thm) *  | 
| 
3409
 
c0466958df5d
Tidying of signature.  More robust renaming in freeze_thaw.
 
paulson 
parents: 
2814 
diff
changeset
 | 
61  | 
((int*(bool*thm))Net.net * (int*(bool*thm))Net.net) ->  | 
| 1801 | 62  | 
(int*(bool*thm))Net.net * (int*(bool*thm))Net.net  | 
| 10805 | 63  | 
val is_fact : thm -> bool  | 
64  | 
val lessb : (bool * thm) * (bool * thm) -> bool  | 
|
65  | 
val lift_inst_rule : thm * int * (string*string)list * thm -> thm  | 
|
66  | 
val make_elim : thm -> thm  | 
|
67  | 
val match_from_net_tac : (int*thm) Net.net -> int -> tactic  | 
|
68  | 
val match_tac : thm list -> int -> tactic  | 
|
69  | 
val metacut_tac : thm -> int -> tactic  | 
|
70  | 
val net_bimatch_tac : (bool*thm) list -> int -> tactic  | 
|
71  | 
val net_biresolve_tac : (bool*thm) list -> int -> tactic  | 
|
72  | 
val net_match_tac : thm list -> int -> tactic  | 
|
73  | 
val net_resolve_tac : thm list -> int -> tactic  | 
|
74  | 
val norm_hhf : thm -> thm  | 
|
75  | 
val norm_hhf_tac : int -> tactic  | 
|
| 10817 | 76  | 
val PRIMITIVE : (thm -> thm) -> tactic  | 
77  | 
val PRIMSEQ : (thm -> thm Seq.seq) -> tactic  | 
|
| 10805 | 78  | 
val prune_params_tac : tactic  | 
79  | 
val rename_params_tac : string list -> int -> tactic  | 
|
80  | 
val rename_tac : string -> int -> tactic  | 
|
81  | 
val rename_last_tac : string -> string list -> int -> tactic  | 
|
82  | 
val resolve_from_net_tac : (int*thm) Net.net -> int -> tactic  | 
|
83  | 
val resolve_tac : thm list -> int -> tactic  | 
|
| 10817 | 84  | 
val res_inst_tac : (string*string)list -> thm -> int -> tactic  | 
| 10444 | 85  | 
val rewrite_goal_tac : thm list -> int -> tactic  | 
| 
3575
 
4e9beacb5339
improved rewrite_thm / rewrite_goals to handle conditional eqns;
 
wenzelm 
parents: 
3554 
diff
changeset
 | 
86  | 
val rewrite_goals_rule: thm list -> thm -> thm  | 
| 10805 | 87  | 
val rewrite_rule : thm list -> thm -> thm  | 
88  | 
val rewrite_goals_tac : thm list -> tactic  | 
|
89  | 
val rewrite_tac : thm list -> tactic  | 
|
90  | 
val rewtac : thm -> tactic  | 
|
91  | 
val rotate_tac : int -> int -> tactic  | 
|
92  | 
val rtac : thm -> int -> tactic  | 
|
93  | 
val rule_by_tactic : tactic -> thm -> thm  | 
|
94  | 
val solve_tac : thm list -> int -> tactic  | 
|
95  | 
val subgoal_tac : string -> int -> tactic  | 
|
96  | 
val subgoals_tac : string list -> int -> tactic  | 
|
97  | 
val subgoals_of_brl : bool * thm -> int  | 
|
98  | 
val term_lift_inst_rule :  | 
|
| 
1975
 
eec67972b1bf
renamed cterm_lift_inst_rule to term_lift_inst_rule and made it take
 
nipkow 
parents: 
1966 
diff
changeset
 | 
99  | 
thm * int * (indexname*typ)list * ((indexname*typ)*term)list * thm  | 
| 
 
eec67972b1bf
renamed cterm_lift_inst_rule to term_lift_inst_rule and made it take
 
nipkow 
parents: 
1966 
diff
changeset
 | 
100  | 
-> thm  | 
| 10347 | 101  | 
val instantiate_tac : (string * string) list -> tactic  | 
| 10805 | 102  | 
val thin_tac : string -> int -> tactic  | 
103  | 
val trace_goalno_tac : (int -> tactic) -> int -> tactic  | 
|
| 11774 | 104  | 
end;  | 
| 0 | 105  | 
|
| 11774 | 106  | 
signature TACTIC =  | 
107  | 
sig  | 
|
108  | 
include BASIC_TACTIC  | 
|
| 11929 | 109  | 
val innermost_params: int -> thm -> (string * typ) list  | 
| 11774 | 110  | 
val untaglist: (int * 'a) list -> 'a list  | 
111  | 
val orderlist: (int * 'a) list -> 'a list  | 
|
112  | 
val rewrite: bool -> thm list -> cterm -> thm  | 
|
113  | 
val rewrite_cterm: bool -> thm list -> cterm -> cterm  | 
|
114  | 
val simplify: bool -> thm list -> thm -> thm  | 
|
| 
11961
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
115  | 
val prove: theory -> string list -> term list -> term -> (thm list -> tactic) -> thm  | 
| 11774 | 116  | 
end;  | 
| 0 | 117  | 
|
| 11774 | 118  | 
structure Tactic: TACTIC =  | 
| 0 | 119  | 
struct  | 
120  | 
||
| 1501 | 121  | 
(*Discover which goal is chosen: SOMEGOAL(trace_goalno_tac tac) *)  | 
| 10817 | 122  | 
fun trace_goalno_tac tac i st =  | 
| 4270 | 123  | 
case Seq.pull(tac i st) of  | 
| 10805 | 124  | 
None => Seq.empty  | 
| 10817 | 125  | 
      | seqcell => (writeln ("Subgoal " ^ string_of_int i ^ " selected");
 | 
| 10805 | 126  | 
Seq.make(fn()=> seqcell));  | 
| 0 | 127  | 
|
128  | 
(*Makes a rule by applying a tactic to an existing rule*)  | 
|
| 1501 | 129  | 
fun rule_by_tactic tac rl =  | 
| 2688 | 130  | 
let val (st, thaw) = freeze_thaw (zero_var_indexes rl)  | 
| 4270 | 131  | 
in case Seq.pull (tac st) of  | 
| 10805 | 132  | 
        None        => raise THM("rule_by_tactic", 0, [rl])
 | 
| 2688 | 133  | 
| Some(st',_) => Thm.varifyT (thaw st')  | 
134  | 
end;  | 
|
| 10817 | 135  | 
|
| 0 | 136  | 
(*** Basic tactics ***)  | 
137  | 
||
138  | 
(*Makes a tactic whose effect on a state is given by thmfun: thm->thm seq.*)  | 
|
| 4270 | 139  | 
fun PRIMSEQ thmfun st = thmfun st handle THM _ => Seq.empty;  | 
| 0 | 140  | 
|
141  | 
(*Makes a tactic whose effect on a state is given by thmfun: thm->thm.*)  | 
|
| 4270 | 142  | 
fun PRIMITIVE thmfun = PRIMSEQ (Seq.single o thmfun);  | 
| 0 | 143  | 
|
144  | 
(*** The following fail if the goal number is out of range:  | 
|
145  | 
thus (REPEAT (resolve_tac rules i)) stops once subgoal i disappears. *)  | 
|
146  | 
||
147  | 
(*Solve subgoal i by assumption*)  | 
|
148  | 
fun assume_tac i = PRIMSEQ (assumption i);  | 
|
149  | 
||
150  | 
(*Solve subgoal i by assumption, using no unification*)  | 
|
151  | 
fun eq_assume_tac i = PRIMITIVE (eq_assumption i);  | 
|
152  | 
||
153  | 
(** Resolution/matching tactics **)  | 
|
154  | 
||
155  | 
(*The composition rule/state: no lifting or var renaming.  | 
|
156  | 
The arg = (bires_flg, orule, m) ; see bicompose for explanation.*)  | 
|
157  | 
fun compose_tac arg i = PRIMSEQ (bicompose false arg i);  | 
|
158  | 
||
159  | 
(*Converts a "destruct" rule like P&Q==>P to an "elimination" rule  | 
|
160  | 
like [| P&Q; P==>R |] ==> R *)  | 
|
161  | 
fun make_elim rl = zero_var_indexes (rl RS revcut_rl);  | 
|
162  | 
||
163  | 
(*Attack subgoal i by resolution, using flags to indicate elimination rules*)  | 
|
164  | 
fun biresolve_tac brules i = PRIMSEQ (biresolution false brules i);  | 
|
165  | 
||
166  | 
(*Resolution: the simple case, works for introduction rules*)  | 
|
167  | 
fun resolve_tac rules = biresolve_tac (map (pair false) rules);  | 
|
168  | 
||
169  | 
(*Resolution with elimination rules only*)  | 
|
170  | 
fun eresolve_tac rules = biresolve_tac (map (pair true) rules);  | 
|
171  | 
||
172  | 
(*Forward reasoning using destruction rules.*)  | 
|
173  | 
fun forward_tac rls = resolve_tac (map make_elim rls) THEN' assume_tac;  | 
|
174  | 
||
175  | 
(*Like forward_tac, but deletes the assumption after use.*)  | 
|
176  | 
fun dresolve_tac rls = eresolve_tac (map make_elim rls);  | 
|
177  | 
||
178  | 
(*Shorthand versions: for resolution with a single theorem*)  | 
|
| 7491 | 179  | 
val atac = assume_tac;  | 
180  | 
fun rtac rl = resolve_tac [rl];  | 
|
181  | 
fun dtac rl = dresolve_tac [rl];  | 
|
| 1460 | 182  | 
fun etac rl = eresolve_tac [rl];  | 
| 7491 | 183  | 
fun ftac rl = forward_tac [rl];  | 
184  | 
fun datac thm j = EVERY' (dtac thm::replicate j atac);  | 
|
185  | 
fun eatac thm j = EVERY' (etac thm::replicate j atac);  | 
|
186  | 
fun fatac thm j = EVERY' (ftac thm::replicate j atac);  | 
|
| 0 | 187  | 
|
188  | 
(*Use an assumption or some rules ... A popular combination!*)  | 
|
189  | 
fun ares_tac rules = assume_tac ORELSE' resolve_tac rules;  | 
|
190  | 
||
| 5263 | 191  | 
fun solve_tac rules = resolve_tac rules THEN_ALL_NEW assume_tac;  | 
192  | 
||
| 0 | 193  | 
(*Matching tactics -- as above, but forbid updating of state*)  | 
194  | 
fun bimatch_tac brules i = PRIMSEQ (biresolution true brules i);  | 
|
195  | 
fun match_tac rules = bimatch_tac (map (pair false) rules);  | 
|
196  | 
fun ematch_tac rules = bimatch_tac (map (pair true) rules);  | 
|
197  | 
fun dmatch_tac rls = ematch_tac (map make_elim rls);  | 
|
198  | 
||
199  | 
(*Smash all flex-flex disagreement pairs in the proof state.*)  | 
|
200  | 
val flexflex_tac = PRIMSEQ flexflex_rule;  | 
|
201  | 
||
| 
3409
 
c0466958df5d
Tidying of signature.  More robust renaming in freeze_thaw.
 
paulson 
parents: 
2814 
diff
changeset
 | 
202  | 
|
| 
 
c0466958df5d
Tidying of signature.  More robust renaming in freeze_thaw.
 
paulson 
parents: 
2814 
diff
changeset
 | 
203  | 
(*Remove duplicate subgoals. By Mark Staples*)  | 
| 
 
c0466958df5d
Tidying of signature.  More robust renaming in freeze_thaw.
 
paulson 
parents: 
2814 
diff
changeset
 | 
204  | 
local  | 
| 
 
c0466958df5d
Tidying of signature.  More robust renaming in freeze_thaw.
 
paulson 
parents: 
2814 
diff
changeset
 | 
205  | 
fun cterm_aconv (a,b) = #t (rep_cterm a) aconv #t (rep_cterm b);  | 
| 
 
c0466958df5d
Tidying of signature.  More robust renaming in freeze_thaw.
 
paulson 
parents: 
2814 
diff
changeset
 | 
206  | 
in  | 
| 10817 | 207  | 
fun distinct_subgoals_tac state =  | 
| 
3409
 
c0466958df5d
Tidying of signature.  More robust renaming in freeze_thaw.
 
paulson 
parents: 
2814 
diff
changeset
 | 
208  | 
let val (frozth,thawfn) = freeze_thaw state  | 
| 10805 | 209  | 
val froz_prems = cprems_of frozth  | 
210  | 
val assumed = implies_elim_list frozth (map assume froz_prems)  | 
|
211  | 
val implied = implies_intr_list (gen_distinct cterm_aconv froz_prems)  | 
|
212  | 
assumed;  | 
|
| 4270 | 213  | 
in Seq.single (thawfn implied) end  | 
| 10817 | 214  | 
end;  | 
| 
3409
 
c0466958df5d
Tidying of signature.  More robust renaming in freeze_thaw.
 
paulson 
parents: 
2814 
diff
changeset
 | 
215  | 
|
| 
 
c0466958df5d
Tidying of signature.  More robust renaming in freeze_thaw.
 
paulson 
parents: 
2814 
diff
changeset
 | 
216  | 
|
| 11929 | 217  | 
(*Determine print names of goal parameters (reversed)*)  | 
218  | 
fun innermost_params i st =  | 
|
219  | 
let val (_, _, Bi, _) = dest_state (st, i)  | 
|
220  | 
in rename_wrt_term Bi (Logic.strip_params Bi) end;  | 
|
221  | 
||
| 0 | 222  | 
(*Lift and instantiate a rule wrt the given state and subgoal number *)  | 
| 1501 | 223  | 
fun lift_inst_rule (st, i, sinsts, rule) =  | 
224  | 
let val {maxidx,sign,...} = rep_thm st
 | 
|
225  | 
val (_, _, Bi, _) = dest_state(st,i)  | 
|
| 10805 | 226  | 
val params = Logic.strip_params Bi (*params of subgoal i*)  | 
| 0 | 227  | 
val params = rev(rename_wrt_term Bi params) (*as they are printed*)  | 
228  | 
val paramTs = map #2 params  | 
|
229  | 
and inc = maxidx+1  | 
|
230  | 
fun liftvar (Var ((a,j), T)) = Var((a, j+inc), paramTs---> incr_tvar inc T)  | 
|
231  | 
      | liftvar t = raise TERM("Variable expected", [t]);
 | 
|
| 10817 | 232  | 
fun liftterm t = list_abs_free (params,  | 
| 10805 | 233  | 
Logic.incr_indexes(paramTs,inc) t)  | 
| 0 | 234  | 
(*Lifts instantiation pair over params*)  | 
| 230 | 235  | 
fun liftpair (cv,ct) = (cterm_fun liftvar cv, cterm_fun liftterm ct)  | 
| 0 | 236  | 
fun lifttvar((a,i),ctyp) =  | 
| 10805 | 237  | 
        let val {T,sign} = rep_ctyp ctyp
 | 
238  | 
in ((a,i+inc), ctyp_of sign (incr_tvar inc T)) end  | 
|
| 1501 | 239  | 
val rts = types_sorts rule and (types,sorts) = types_sorts st  | 
| 0 | 240  | 
fun types'(a,~1) = (case assoc(params,a) of None => types(a,~1) | sm => sm)  | 
241  | 
| types'(ixn) = types ixn;  | 
|
| 
949
 
83c588d6fee9
Changed treatment of during type inference internally generated type
 
nipkow 
parents: 
947 
diff
changeset
 | 
242  | 
val used = add_term_tvarnames  | 
| 1501 | 243  | 
(#prop(rep_thm st) $ #prop(rep_thm rule),[])  | 
| 
949
 
83c588d6fee9
Changed treatment of during type inference internally generated type
 
nipkow 
parents: 
947 
diff
changeset
 | 
244  | 
val (Tinsts,insts) = read_insts sign rts (types',sorts) used sinsts  | 
| 
8129
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
7596 
diff
changeset
 | 
245  | 
in Drule.instantiate (map lifttvar Tinsts, map liftpair insts)  | 
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
7596 
diff
changeset
 | 
246  | 
(lift_rule (st,i) rule)  | 
| 0 | 247  | 
end;  | 
248  | 
||
| 3984 | 249  | 
(*  | 
250  | 
Like lift_inst_rule but takes terms, not strings, where the terms may contain  | 
|
251  | 
Bounds referring to parameters of the subgoal.  | 
|
252  | 
||
253  | 
insts: [...,(vj,tj),...]  | 
|
254  | 
||
255  | 
The tj may contain references to parameters of subgoal i of the state st  | 
|
256  | 
in the form of Bound k, i.e. the tj may be subterms of the subgoal.  | 
|
257  | 
To saturate the lose bound vars, the tj are enclosed in abstractions  | 
|
258  | 
corresponding to the parameters of subgoal i, thus turning them into  | 
|
259  | 
functions. At the same time, the types of the vj are lifted.  | 
|
260  | 
||
261  | 
NB: the types in insts must be correctly instantiated already,  | 
|
262  | 
i.e. Tinsts is not applied to insts.  | 
|
263  | 
*)  | 
|
| 
1975
 
eec67972b1bf
renamed cterm_lift_inst_rule to term_lift_inst_rule and made it take
 
nipkow 
parents: 
1966 
diff
changeset
 | 
264  | 
fun term_lift_inst_rule (st, i, Tinsts, insts, rule) =  | 
| 1966 | 265  | 
let val {maxidx,sign,...} = rep_thm st
 | 
266  | 
val (_, _, Bi, _) = dest_state(st,i)  | 
|
267  | 
val params = Logic.strip_params Bi (*params of subgoal i*)  | 
|
268  | 
val paramTs = map #2 params  | 
|
269  | 
and inc = maxidx+1  | 
|
| 
1975
 
eec67972b1bf
renamed cterm_lift_inst_rule to term_lift_inst_rule and made it take
 
nipkow 
parents: 
1966 
diff
changeset
 | 
270  | 
fun liftvar ((a,j), T) = Var((a, j+inc), paramTs---> incr_tvar inc T)  | 
| 
 
eec67972b1bf
renamed cterm_lift_inst_rule to term_lift_inst_rule and made it take
 
nipkow 
parents: 
1966 
diff
changeset
 | 
271  | 
(*lift only Var, not term, which must be lifted already*)  | 
| 
 
eec67972b1bf
renamed cterm_lift_inst_rule to term_lift_inst_rule and made it take
 
nipkow 
parents: 
1966 
diff
changeset
 | 
272  | 
fun liftpair (v,t) = (cterm_of sign (liftvar v), cterm_of sign t)  | 
| 
 
eec67972b1bf
renamed cterm_lift_inst_rule to term_lift_inst_rule and made it take
 
nipkow 
parents: 
1966 
diff
changeset
 | 
273  | 
fun liftTpair((a,i),T) = ((a,i+inc), ctyp_of sign (incr_tvar inc T))  | 
| 
8129
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
7596 
diff
changeset
 | 
274  | 
in Drule.instantiate (map liftTpair Tinsts, map liftpair insts)  | 
| 
 
29e239c7b8c2
Thm.instantiate no longer normalizes, but Drule.instantiate does
 
paulson 
parents: 
7596 
diff
changeset
 | 
275  | 
(lift_rule (st,i) rule)  | 
| 1966 | 276  | 
end;  | 
| 0 | 277  | 
|
278  | 
(*** Resolve after lifting and instantation; may refer to parameters of the  | 
|
279  | 
subgoal. Fails if "i" is out of range. ***)  | 
|
280  | 
||
281  | 
(*compose version: arguments are as for bicompose.*)  | 
|
| 10817 | 282  | 
fun compose_inst_tac sinsts (bires_flg, rule, nsubgoal) i st =  | 
| 
8977
 
dd8bc754a072
res_inst_tac, etc., no longer print the "dest_state" message when the selected
 
paulson 
parents: 
8129 
diff
changeset
 | 
283  | 
if i > nprems_of st then no_tac st  | 
| 
 
dd8bc754a072
res_inst_tac, etc., no longer print the "dest_state" message when the selected
 
paulson 
parents: 
8129 
diff
changeset
 | 
284  | 
else st |>  | 
| 
 
dd8bc754a072
res_inst_tac, etc., no longer print the "dest_state" message when the selected
 
paulson 
parents: 
8129 
diff
changeset
 | 
285  | 
(compose_tac (bires_flg, lift_inst_rule (st, i, sinsts, rule), nsubgoal) i  | 
| 
 
dd8bc754a072
res_inst_tac, etc., no longer print the "dest_state" message when the selected
 
paulson 
parents: 
8129 
diff
changeset
 | 
286  | 
handle TERM (msg,_) => (writeln msg; no_tac)  | 
| 10805 | 287  | 
| THM (msg,_,_) => (writeln msg; no_tac));  | 
| 0 | 288  | 
|
| 761 | 289  | 
(*"Resolve" version. Note: res_inst_tac cannot behave sensibly if the  | 
290  | 
terms that are substituted contain (term or type) unknowns from the  | 
|
291  | 
goal, because it is unable to instantiate goal unknowns at the same time.  | 
|
292  | 
||
| 2029 | 293  | 
The type checker is instructed not to freeze flexible type vars that  | 
| 
952
 
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
 
nipkow 
parents: 
949 
diff
changeset
 | 
294  | 
were introduced during type inference and still remain in the term at the  | 
| 
 
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
 
nipkow 
parents: 
949 
diff
changeset
 | 
295  | 
end. This increases flexibility but can introduce schematic type vars in  | 
| 
 
9e10962866b0
Removed an old bug which made some simultaneous instantiations fail if they
 
nipkow 
parents: 
949 
diff
changeset
 | 
296  | 
goals.  | 
| 761 | 297  | 
*)  | 
| 0 | 298  | 
fun res_inst_tac sinsts rule i =  | 
299  | 
compose_inst_tac sinsts (false, rule, nprems_of rule) i;  | 
|
300  | 
||
| 1501 | 301  | 
(*eresolve elimination version*)  | 
| 0 | 302  | 
fun eres_inst_tac sinsts rule i =  | 
303  | 
compose_inst_tac sinsts (true, rule, nprems_of rule) i;  | 
|
304  | 
||
| 
270
 
d506ea00c825
tactic/make_elim_preserve: recoded to avoid using lift_inst_rule.  Instead
 
lcp 
parents: 
230 
diff
changeset
 | 
305  | 
(*For forw_inst_tac and dres_inst_tac. Preserve Var indexes of rl;  | 
| 
 
d506ea00c825
tactic/make_elim_preserve: recoded to avoid using lift_inst_rule.  Instead
 
lcp 
parents: 
230 
diff
changeset
 | 
306  | 
increment revcut_rl instead.*)  | 
| 10817 | 307  | 
fun make_elim_preserve rl =  | 
| 
270
 
d506ea00c825
tactic/make_elim_preserve: recoded to avoid using lift_inst_rule.  Instead
 
lcp 
parents: 
230 
diff
changeset
 | 
308  | 
  let val {maxidx,...} = rep_thm rl
 | 
| 6390 | 309  | 
fun cvar ixn = cterm_of (Theory.sign_of ProtoPure.thy) (Var(ixn,propT));  | 
| 10817 | 310  | 
val revcut_rl' =  | 
| 10805 | 311  | 
          instantiate ([],  [(cvar("V",0), cvar("V",maxidx+1)),
 | 
312  | 
                             (cvar("W",0), cvar("W",maxidx+1))]) revcut_rl
 | 
|
| 0 | 313  | 
val arg = (false, rl, nprems_of rl)  | 
| 4270 | 314  | 
val [th] = Seq.list_of (bicompose false arg 1 revcut_rl')  | 
| 0 | 315  | 
in th end  | 
316  | 
  handle Bind => raise THM("make_elim_preserve", 1, [rl]);
 | 
|
317  | 
||
| 
270
 
d506ea00c825
tactic/make_elim_preserve: recoded to avoid using lift_inst_rule.  Instead
 
lcp 
parents: 
230 
diff
changeset
 | 
318  | 
(*instantiate and cut -- for a FACT, anyway...*)  | 
| 
 
d506ea00c825
tactic/make_elim_preserve: recoded to avoid using lift_inst_rule.  Instead
 
lcp 
parents: 
230 
diff
changeset
 | 
319  | 
fun cut_inst_tac sinsts rule = res_inst_tac sinsts (make_elim_preserve rule);  | 
| 0 | 320  | 
|
| 
270
 
d506ea00c825
tactic/make_elim_preserve: recoded to avoid using lift_inst_rule.  Instead
 
lcp 
parents: 
230 
diff
changeset
 | 
321  | 
(*forward tactic applies a RULE to an assumption without deleting it*)  | 
| 
 
d506ea00c825
tactic/make_elim_preserve: recoded to avoid using lift_inst_rule.  Instead
 
lcp 
parents: 
230 
diff
changeset
 | 
322  | 
fun forw_inst_tac sinsts rule = cut_inst_tac sinsts rule THEN' assume_tac;  | 
| 
 
d506ea00c825
tactic/make_elim_preserve: recoded to avoid using lift_inst_rule.  Instead
 
lcp 
parents: 
230 
diff
changeset
 | 
323  | 
|
| 
 
d506ea00c825
tactic/make_elim_preserve: recoded to avoid using lift_inst_rule.  Instead
 
lcp 
parents: 
230 
diff
changeset
 | 
324  | 
(*dresolve tactic applies a RULE to replace an assumption*)  | 
| 0 | 325  | 
fun dres_inst_tac sinsts rule = eres_inst_tac sinsts (make_elim_preserve rule);  | 
326  | 
||
| 10347 | 327  | 
(*instantiate variables in the whole state*)  | 
328  | 
val instantiate_tac = PRIMITIVE o read_instantiate;  | 
|
329  | 
||
| 1951 | 330  | 
(*Deletion of an assumption*)  | 
331  | 
fun thin_tac s = eres_inst_tac [("V",s)] thin_rl;
 | 
|
332  | 
||
| 
270
 
d506ea00c825
tactic/make_elim_preserve: recoded to avoid using lift_inst_rule.  Instead
 
lcp 
parents: 
230 
diff
changeset
 | 
333  | 
(*** Applications of cut_rl ***)  | 
| 0 | 334  | 
|
335  | 
(*Used by metacut_tac*)  | 
|
336  | 
fun bires_cut_tac arg i =  | 
|
| 1460 | 337  | 
resolve_tac [cut_rl] i THEN biresolve_tac arg (i+1) ;  | 
| 0 | 338  | 
|
339  | 
(*The conclusion of the rule gets assumed in subgoal i,  | 
|
340  | 
while subgoal i+1,... are the premises of the rule.*)  | 
|
341  | 
fun metacut_tac rule = bires_cut_tac [(false,rule)];  | 
|
342  | 
||
343  | 
(*Recognizes theorems that are not rules, but simple propositions*)  | 
|
344  | 
fun is_fact rl =  | 
|
345  | 
case prems_of rl of  | 
|
| 10805 | 346  | 
[] => true | _::_ => false;  | 
| 0 | 347  | 
|
348  | 
(*"Cut" all facts from theorem list into the goal as assumptions. *)  | 
|
349  | 
fun cut_facts_tac ths i =  | 
|
350  | 
EVERY (map (fn th => metacut_tac th i) (filter is_fact ths));  | 
|
351  | 
||
352  | 
(*Introduce the given proposition as a lemma and subgoal*)  | 
|
| 10817 | 353  | 
fun subgoal_tac sprop i st =  | 
| 4270 | 354  | 
  let val st'    = Seq.hd (res_inst_tac [("psi", sprop)] cut_rl i st)
 | 
| 
4178
 
e64ff1c1bc70
subgoal_tac displays a warning if the new subgoal has type variables
 
paulson 
parents: 
3991 
diff
changeset
 | 
355  | 
val concl' = Logic.strip_assums_concl (List.nth(prems_of st', i))  | 
| 10817 | 356  | 
in  | 
| 
4178
 
e64ff1c1bc70
subgoal_tac displays a warning if the new subgoal has type variables
 
paulson 
parents: 
3991 
diff
changeset
 | 
357  | 
if null (term_tvars concl') then ()  | 
| 
 
e64ff1c1bc70
subgoal_tac displays a warning if the new subgoal has type variables
 
paulson 
parents: 
3991 
diff
changeset
 | 
358  | 
else warning"Type variables in new subgoal: add a type constraint?";  | 
| 4270 | 359  | 
Seq.single st'  | 
| 
4178
 
e64ff1c1bc70
subgoal_tac displays a warning if the new subgoal has type variables
 
paulson 
parents: 
3991 
diff
changeset
 | 
360  | 
end;  | 
| 0 | 361  | 
|
| 439 | 362  | 
(*Introduce a list of lemmas and subgoals*)  | 
363  | 
fun subgoals_tac sprops = EVERY' (map subgoal_tac sprops);  | 
|
364  | 
||
| 0 | 365  | 
|
366  | 
(**** Indexing and filtering of theorems ****)  | 
|
367  | 
||
368  | 
(*Returns the list of potentially resolvable theorems for the goal "prem",  | 
|
| 10805 | 369  | 
using the predicate could(subgoal,concl).  | 
| 0 | 370  | 
Resulting list is no longer than "limit"*)  | 
371  | 
fun filter_thms could (limit, prem, ths) =  | 
|
372  | 
let val pb = Logic.strip_assums_concl prem; (*delete assumptions*)  | 
|
373  | 
fun filtr (limit, []) = []  | 
|
| 10805 | 374  | 
| filtr (limit, th::ths) =  | 
375  | 
if limit=0 then []  | 
|
376  | 
else if could(pb, concl_of th) then th :: filtr(limit-1, ths)  | 
|
377  | 
else filtr(limit,ths)  | 
|
| 0 | 378  | 
in filtr(limit,ths) end;  | 
379  | 
||
380  | 
||
381  | 
(*** biresolution and resolution using nets ***)  | 
|
382  | 
||
383  | 
(** To preserve the order of the rules, tag them with increasing integers **)  | 
|
384  | 
||
385  | 
(*insert tags*)  | 
|
386  | 
fun taglist k [] = []  | 
|
387  | 
| taglist k (x::xs) = (k,x) :: taglist (k+1) xs;  | 
|
388  | 
||
389  | 
(*remove tags and suppress duplicates -- list is assumed sorted!*)  | 
|
390  | 
fun untaglist [] = []  | 
|
391  | 
| untaglist [(k:int,x)] = [x]  | 
|
392  | 
| untaglist ((k,x) :: (rest as (k',x')::_)) =  | 
|
393  | 
if k=k' then untaglist rest  | 
|
394  | 
else x :: untaglist rest;  | 
|
395  | 
||
396  | 
(*return list elements in original order*)  | 
|
| 10817 | 397  | 
fun orderlist kbrls = untaglist (sort (int_ord o pairself fst) kbrls);  | 
| 0 | 398  | 
|
399  | 
(*insert one tagged brl into the pair of nets*)  | 
|
| 
1077
 
c2df11ae8b55
Renamed insert_kbrl to insert_tagged_brl and exported it.  Now
 
lcp 
parents: 
952 
diff
changeset
 | 
400  | 
fun insert_tagged_brl (kbrl as (k,(eres,th)), (inet,enet)) =  | 
| 10817 | 401  | 
if eres then  | 
| 10805 | 402  | 
case prems_of th of  | 
403  | 
prem::_ => (inet, Net.insert_term ((prem,kbrl), enet, K false))  | 
|
404  | 
| [] => error"insert_tagged_brl: elimination rule with no premises"  | 
|
| 0 | 405  | 
else (Net.insert_term ((concl_of th, kbrl), inet, K false), enet);  | 
406  | 
||
407  | 
(*build a pair of nets for biresolution*)  | 
|
| 10817 | 408  | 
fun build_netpair netpair brls =  | 
| 
1077
 
c2df11ae8b55
Renamed insert_kbrl to insert_tagged_brl and exported it.  Now
 
lcp 
parents: 
952 
diff
changeset
 | 
409  | 
foldr insert_tagged_brl (taglist 1 brls, netpair);  | 
| 0 | 410  | 
|
| 1801 | 411  | 
(*delete one kbrl from the pair of nets;  | 
412  | 
we don't know the value of k, so we use 0 and ignore it in the comparison*)  | 
|
413  | 
local  | 
|
414  | 
fun eq_kbrl ((k,(eres,th)), (k',(eres',th'))) = eq_thm (th,th')  | 
|
415  | 
in  | 
|
416  | 
fun delete_tagged_brl (brl as (eres,th), (inet,enet)) =  | 
|
| 10817 | 417  | 
if eres then  | 
| 10805 | 418  | 
case prems_of th of  | 
419  | 
prem::_ => (inet, Net.delete_term ((prem, (0,brl)), enet, eq_kbrl))  | 
|
420  | 
| [] => (inet,enet) (*no major premise: ignore*)  | 
|
| 1801 | 421  | 
else (Net.delete_term ((concl_of th, (0,brl)), inet, eq_kbrl), enet);  | 
422  | 
end;  | 
|
423  | 
||
424  | 
||
| 10817 | 425  | 
(*biresolution using a pair of nets rather than rules.  | 
| 
3706
 
e57b5902822f
Generalized and exported biresolution_from_nets_tac to allow the declaration
 
paulson 
parents: 
3575 
diff
changeset
 | 
426  | 
function "order" must sort and possibly filter the list of brls.  | 
| 
 
e57b5902822f
Generalized and exported biresolution_from_nets_tac to allow the declaration
 
paulson 
parents: 
3575 
diff
changeset
 | 
427  | 
boolean "match" indicates matching or unification.*)  | 
| 
 
e57b5902822f
Generalized and exported biresolution_from_nets_tac to allow the declaration
 
paulson 
parents: 
3575 
diff
changeset
 | 
428  | 
fun biresolution_from_nets_tac order match (inet,enet) =  | 
| 0 | 429  | 
SUBGOAL  | 
430  | 
(fn (prem,i) =>  | 
|
431  | 
let val hyps = Logic.strip_assums_hyp prem  | 
|
| 10817 | 432  | 
and concl = Logic.strip_assums_concl prem  | 
| 0 | 433  | 
val kbrls = Net.unify_term inet concl @  | 
| 
2672
 
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
 
paulson 
parents: 
2580 
diff
changeset
 | 
434  | 
List.concat (map (Net.unify_term enet) hyps)  | 
| 
3706
 
e57b5902822f
Generalized and exported biresolution_from_nets_tac to allow the declaration
 
paulson 
parents: 
3575 
diff
changeset
 | 
435  | 
in PRIMSEQ (biresolution match (order kbrls) i) end);  | 
| 0 | 436  | 
|
| 
3706
 
e57b5902822f
Generalized and exported biresolution_from_nets_tac to allow the declaration
 
paulson 
parents: 
3575 
diff
changeset
 | 
437  | 
(*versions taking pre-built nets. No filtering of brls*)  | 
| 
 
e57b5902822f
Generalized and exported biresolution_from_nets_tac to allow the declaration
 
paulson 
parents: 
3575 
diff
changeset
 | 
438  | 
val biresolve_from_nets_tac = biresolution_from_nets_tac orderlist false;  | 
| 
 
e57b5902822f
Generalized and exported biresolution_from_nets_tac to allow the declaration
 
paulson 
parents: 
3575 
diff
changeset
 | 
439  | 
val bimatch_from_nets_tac = biresolution_from_nets_tac orderlist true;  | 
| 0 | 440  | 
|
441  | 
(*fast versions using nets internally*)  | 
|
| 670 | 442  | 
val net_biresolve_tac =  | 
443  | 
biresolve_from_nets_tac o build_netpair(Net.empty,Net.empty);  | 
|
444  | 
||
445  | 
val net_bimatch_tac =  | 
|
446  | 
bimatch_from_nets_tac o build_netpair(Net.empty,Net.empty);  | 
|
| 0 | 447  | 
|
448  | 
(*** Simpler version for resolve_tac -- only one net, and no hyps ***)  | 
|
449  | 
||
450  | 
(*insert one tagged rl into the net*)  | 
|
451  | 
fun insert_krl (krl as (k,th), net) =  | 
|
452  | 
Net.insert_term ((concl_of th, krl), net, K false);  | 
|
453  | 
||
454  | 
(*build a net of rules for resolution*)  | 
|
| 10817 | 455  | 
fun build_net rls =  | 
| 0 | 456  | 
foldr insert_krl (taglist 1 rls, Net.empty);  | 
457  | 
||
458  | 
(*resolution using a net rather than rules; pred supports filt_resolve_tac*)  | 
|
459  | 
fun filt_resolution_from_net_tac match pred net =  | 
|
460  | 
SUBGOAL  | 
|
461  | 
(fn (prem,i) =>  | 
|
462  | 
let val krls = Net.unify_term net (Logic.strip_assums_concl prem)  | 
|
| 10817 | 463  | 
in  | 
464  | 
if pred krls  | 
|
| 0 | 465  | 
then PRIMSEQ  | 
| 10805 | 466  | 
(biresolution match (map (pair false) (orderlist krls)) i)  | 
| 0 | 467  | 
else no_tac  | 
468  | 
end);  | 
|
469  | 
||
470  | 
(*Resolve the subgoal using the rules (making a net) unless too flexible,  | 
|
471  | 
which means more than maxr rules are unifiable. *)  | 
|
| 10817 | 472  | 
fun filt_resolve_tac rules maxr =  | 
| 0 | 473  | 
let fun pred krls = length krls <= maxr  | 
474  | 
in filt_resolution_from_net_tac false pred (build_net rules) end;  | 
|
475  | 
||
476  | 
(*versions taking pre-built nets*)  | 
|
477  | 
val resolve_from_net_tac = filt_resolution_from_net_tac false (K true);  | 
|
478  | 
val match_from_net_tac = filt_resolution_from_net_tac true (K true);  | 
|
479  | 
||
480  | 
(*fast versions using nets internally*)  | 
|
481  | 
val net_resolve_tac = resolve_from_net_tac o build_net;  | 
|
482  | 
val net_match_tac = match_from_net_tac o build_net;  | 
|
483  | 
||
484  | 
||
485  | 
(*** For Natural Deduction using (bires_flg, rule) pairs ***)  | 
|
486  | 
||
487  | 
(*The number of new subgoals produced by the brule*)  | 
|
| 
1077
 
c2df11ae8b55
Renamed insert_kbrl to insert_tagged_brl and exported it.  Now
 
lcp 
parents: 
952 
diff
changeset
 | 
488  | 
fun subgoals_of_brl (true,rule) = nprems_of rule - 1  | 
| 
 
c2df11ae8b55
Renamed insert_kbrl to insert_tagged_brl and exported it.  Now
 
lcp 
parents: 
952 
diff
changeset
 | 
489  | 
| subgoals_of_brl (false,rule) = nprems_of rule;  | 
| 0 | 490  | 
|
491  | 
(*Less-than test: for sorting to minimize number of new subgoals*)  | 
|
492  | 
fun lessb (brl1,brl2) = subgoals_of_brl brl1 < subgoals_of_brl brl2;  | 
|
493  | 
||
494  | 
||
495  | 
(*** Meta-Rewriting Tactics ***)  | 
|
496  | 
||
497  | 
fun result1 tacf mss thm =  | 
|
| 4270 | 498  | 
apsome fst (Seq.pull (tacf mss thm));  | 
| 0 | 499  | 
|
| 
3575
 
4e9beacb5339
improved rewrite_thm / rewrite_goals to handle conditional eqns;
 
wenzelm 
parents: 
3554 
diff
changeset
 | 
500  | 
val simple_prover =  | 
| 11671 | 501  | 
result1 (fn mss => ALLGOALS (resolve_tac (MetaSimplifier.prems_of_mss mss)));  | 
| 
3575
 
4e9beacb5339
improved rewrite_thm / rewrite_goals to handle conditional eqns;
 
wenzelm 
parents: 
3554 
diff
changeset
 | 
502  | 
|
| 
11768
 
48bc55f43774
unified rewrite/rewrite_cterm/simplify interface;
 
wenzelm 
parents: 
11762 
diff
changeset
 | 
503  | 
val rewrite = MetaSimplifier.rewrite_aux simple_prover;  | 
| 
 
48bc55f43774
unified rewrite/rewrite_cterm/simplify interface;
 
wenzelm 
parents: 
11762 
diff
changeset
 | 
504  | 
val rewrite_cterm = (#2 o Thm.dest_comb o #prop o Thm.crep_thm) ooo rewrite;  | 
| 
 
48bc55f43774
unified rewrite/rewrite_cterm/simplify interface;
 
wenzelm 
parents: 
11762 
diff
changeset
 | 
505  | 
val simplify = MetaSimplifier.simplify_aux simple_prover;  | 
| 
 
48bc55f43774
unified rewrite/rewrite_cterm/simplify interface;
 
wenzelm 
parents: 
11762 
diff
changeset
 | 
506  | 
val rewrite_rule = simplify true;  | 
| 
10415
 
e6d7b77a0574
moved rewriting functions from Drule to MetaSimplifier
 
berghofe 
parents: 
10347 
diff
changeset
 | 
507  | 
val rewrite_goals_rule = MetaSimplifier.rewrite_goals_rule_aux simple_prover;  | 
| 
3575
 
4e9beacb5339
improved rewrite_thm / rewrite_goals to handle conditional eqns;
 
wenzelm 
parents: 
3554 
diff
changeset
 | 
508  | 
|
| 2145 | 509  | 
(*Rewrite subgoal i only. SELECT_GOAL avoids inefficiencies in goals_conv.*)  | 
510  | 
fun asm_rewrite_goal_tac mode prover_tac mss =  | 
|
| 11671 | 511  | 
SELECT_GOAL  | 
512  | 
(PRIMITIVE (MetaSimplifier.rewrite_goal_rule mode (result1 prover_tac) mss 1));  | 
|
| 0 | 513  | 
|
| 10444 | 514  | 
fun rewrite_goal_tac rews =  | 
515  | 
asm_rewrite_goal_tac (true, false, false) (K no_tac) (MetaSimplifier.mss_of rews);  | 
|
516  | 
||
| 
69
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
517  | 
(*Rewrite throughout proof state. *)  | 
| 
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
518  | 
fun rewrite_tac defs = PRIMITIVE(rewrite_rule defs);  | 
| 0 | 519  | 
|
520  | 
(*Rewrite subgoals only, not main goal. *)  | 
|
| 
69
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
521  | 
fun rewrite_goals_tac defs = PRIMITIVE (rewrite_goals_rule defs);  | 
| 1460 | 522  | 
fun rewtac def = rewrite_goals_tac [def];  | 
| 0 | 523  | 
|
| 10817 | 524  | 
fun norm_hhf th =  | 
525  | 
(if Logic.is_norm_hhf (#prop (Thm.rep_thm th)) then th else rewrite_rule [Drule.norm_hhf_eq] th)  | 
|
526  | 
|> Drule.forall_elim_vars_safe;  | 
|
527  | 
||
528  | 
val norm_hhf_tac = SUBGOAL (fn (t, i) =>  | 
|
529  | 
if Logic.is_norm_hhf t then all_tac  | 
|
530  | 
else rewrite_goal_tac [Drule.norm_hhf_eq] i);  | 
|
| 10805 | 531  | 
|
| 0 | 532  | 
|
| 1501 | 533  | 
(*** for folding definitions, handling critical pairs ***)  | 
| 
69
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
534  | 
|
| 
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
535  | 
(*The depth of nesting in a term*)  | 
| 
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
536  | 
fun term_depth (Abs(a,T,t)) = 1 + term_depth t  | 
| 2145 | 537  | 
| term_depth (f$t) = 1 + Int.max(term_depth f, term_depth t)  | 
| 
69
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
538  | 
| term_depth _ = 0;  | 
| 
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
539  | 
|
| 
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
540  | 
val lhs_of_thm = #1 o Logic.dest_equals o #prop o rep_thm;  | 
| 
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
541  | 
|
| 
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
542  | 
(*folding should handle critical pairs! E.g. K == Inl(0), S == Inr(Inl(0))  | 
| 
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
543  | 
Returns longest lhs first to avoid folding its subexpressions.*)  | 
| 
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
544  | 
fun sort_lhs_depths defs =  | 
| 
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
545  | 
let val keylist = make_keylist (term_depth o lhs_of_thm) defs  | 
| 4438 | 546  | 
val keys = distinct (sort (rev_order o int_ord) (map #2 keylist))  | 
| 
69
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
547  | 
in map (keyfilter keylist) keys end;  | 
| 
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
548  | 
|
| 7596 | 549  | 
val rev_defs = sort_lhs_depths o map symmetric;  | 
| 
69
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
550  | 
|
| 7596 | 551  | 
fun fold_rule defs thm = foldl (fn (th, ds) => rewrite_rule ds th) (thm, rev_defs defs);  | 
552  | 
fun fold_tac defs = EVERY (map rewrite_tac (rev_defs defs));  | 
|
553  | 
fun fold_goals_tac defs = EVERY (map rewrite_goals_tac (rev_defs defs));  | 
|
| 
69
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
554  | 
|
| 
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
555  | 
|
| 
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
556  | 
(*** Renaming of parameters in a subgoal  | 
| 
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
557  | 
Names may contain letters, digits or primes and must be  | 
| 
 
e7588b53d6b0
tactic/fold_tac,fold_goals_tac: now handle critical pairs on the
 
lcp 
parents: 
0 
diff
changeset
 | 
558  | 
separated by blanks ***)  | 
| 0 | 559  | 
|
560  | 
(*Calling this will generate the warning "Same as previous level" since  | 
|
561  | 
it affects nothing but the names of bound variables!*)  | 
|
| 9535 | 562  | 
fun rename_params_tac xs i =  | 
563  | 
(if !Logic.auto_rename  | 
|
| 10817 | 564  | 
then (warning "Resetting Logic.auto_rename";  | 
| 10805 | 565  | 
Logic.auto_rename := false)  | 
| 9535 | 566  | 
else (); PRIMITIVE (rename_params_rule (xs, i)));  | 
567  | 
||
| 10817 | 568  | 
fun rename_tac str i =  | 
569  | 
let val cs = Symbol.explode str in  | 
|
| 4693 | 570  | 
case #2 (take_prefix (Symbol.is_letdig orf Symbol.is_blank) cs) of  | 
| 9535 | 571  | 
[] => rename_params_tac (scanwords Symbol.is_letdig cs) i  | 
| 0 | 572  | 
    | c::_ => error ("Illegal character: " ^ c)
 | 
573  | 
end;  | 
|
574  | 
||
| 1501 | 575  | 
(*Rename recent parameters using names generated from a and the suffixes,  | 
576  | 
provided the string a, which represents a term, is an identifier. *)  | 
|
| 10817 | 577  | 
fun rename_last_tac a sufs i =  | 
| 0 | 578  | 
let val names = map (curry op^ a) sufs  | 
579  | 
in if Syntax.is_identifier a  | 
|
580  | 
then PRIMITIVE (rename_params_rule (names,i))  | 
|
581  | 
else all_tac  | 
|
582  | 
end;  | 
|
583  | 
||
| 2043 | 584  | 
(*Prunes all redundant parameters from the proof state by rewriting.  | 
585  | 
DOES NOT rewrite main goal, where quantification over an unused bound  | 
|
586  | 
variable is sometimes done to avoid the need for cut_facts_tac.*)  | 
|
587  | 
val prune_params_tac = rewrite_goals_tac [triv_forall_equality];  | 
|
| 0 | 588  | 
|
| 1501 | 589  | 
(*rotate_tac n i: rotate the assumptions of subgoal i by n positions, from  | 
590  | 
right to left if n is positive, and from left to right if n is negative.*)  | 
|
| 
2672
 
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
 
paulson 
parents: 
2580 
diff
changeset
 | 
591  | 
fun rotate_tac 0 i = all_tac  | 
| 
 
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
 
paulson 
parents: 
2580 
diff
changeset
 | 
592  | 
| rotate_tac k i = PRIMITIVE (rotate_rule k i);  | 
| 1209 | 593  | 
|
| 
7248
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
6979 
diff
changeset
 | 
594  | 
(*Rotates the given subgoal to be the last.*)  | 
| 
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
6979 
diff
changeset
 | 
595  | 
fun defer_tac i = PRIMITIVE (permute_prems (i-1) 1);  | 
| 
 
322151fe6f02
new primitive rule permute_prems to underlie defer_tac and rotate_prems
 
paulson 
parents: 
6979 
diff
changeset
 | 
596  | 
|
| 5974 | 597  | 
(* remove premises that do not satisfy p; fails if all prems satisfy p *)  | 
598  | 
fun filter_prems_tac p =  | 
|
599  | 
let fun Then None tac = Some tac  | 
|
600  | 
| Then (Some tac) tac' = Some(tac THEN' tac');  | 
|
601  | 
fun thins ((tac,n),H) =  | 
|
602  | 
if p H then (tac,n+1)  | 
|
603  | 
else (Then tac (rotate_tac n THEN' etac thin_rl),0);  | 
|
604  | 
in SUBGOAL(fn (subg,n) =>  | 
|
605  | 
let val Hs = Logic.strip_assums_hyp subg  | 
|
606  | 
in case fst(foldl thins ((None,0),Hs)) of  | 
|
607  | 
None => no_tac | Some tac => tac n  | 
|
608  | 
end)  | 
|
609  | 
end;  | 
|
610  | 
||
| 
11961
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
611  | 
|
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
612  | 
(** primitive goal interface for internal use -- avoids "standard" operation *)  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
613  | 
|
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
614  | 
fun prove thy xs asms prop tac =  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
615  | 
let  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
616  | 
val sg = Theory.sign_of thy;  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
617  | 
val statement = Logic.list_implies (asms, prop);  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
618  | 
val frees = map Term.dest_Free (Term.term_frees statement);  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
619  | 
val params = mapfilter (fn x => apsome (pair x) (assoc_string (frees, x))) xs;  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
620  | 
|
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
621  | 
fun err msg = error (msg ^ "\nThe error(s) above occurred for the goal statement:\n" ^  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
622  | 
Sign.string_of_term sg (Term.list_all_free (params, statement)));  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
623  | 
|
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
624  | 
fun cert_safe t = Thm.cterm_of sg t  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
625  | 
handle TERM (msg, _) => err msg | TYPE (msg, _, _) => err msg;  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
626  | 
|
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
627  | 
val _ = cert_safe statement;  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
628  | 
val _ = Term.no_dummy_patterns statement handle TERM (msg, _) => error msg;  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
629  | 
|
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
630  | 
val casms = map cert_safe asms;  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
631  | 
val prems = map (norm_hhf o Thm.assume) casms;  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
632  | 
val goal = Drule.mk_triv_goal (cert_safe prop);  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
633  | 
|
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
634  | 
val goal' =  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
635  | 
(case Seq.pull (tac prems goal) of Some (goal', _) => goal' | _ => err "Tactic failed.");  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
636  | 
val ngoals = Thm.nprems_of goal';  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
637  | 
val raw_result = goal' RS Drule.rev_triv_goal;  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
638  | 
val prop' = #prop (Thm.rep_thm raw_result);  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
639  | 
in  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
640  | 
if ngoals <> 0 then  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
641  | 
      err ("Proof failed.\n" ^ Pretty.string_of (Pretty.chunks (Display.pretty_goals ngoals goal'))
 | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
642  | 
        ^ ("\n" ^ string_of_int ngoals ^ " unsolved goal(s)!"))
 | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
643  | 
else if not (Pattern.matches (Sign.tsig_of sg) (prop, prop')) then  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
644  | 
      err ("Proved a different theorem: " ^ Sign.string_of_term sg prop')
 | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
645  | 
else  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
646  | 
raw_result  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
647  | 
|> Drule.implies_intr_list casms  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
648  | 
|> Drule.forall_intr_list (map (cert_safe o Free) params)  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
649  | 
|> norm_hhf  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
650  | 
|> Thm.varifyT (* FIXME proper scope for polymorphisms!? *)  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
651  | 
end;  | 
| 
 
e5911a25d094
prove: primitive goal interface for internal use;
 
wenzelm 
parents: 
11929 
diff
changeset
 | 
652  | 
|
| 0 | 653  | 
end;  | 
| 1501 | 654  | 
|
| 11774 | 655  | 
structure BasicTactic: BASIC_TACTIC = Tactic;  | 
656  | 
open BasicTactic;  |