1(*  Title:      Provers/splitter.ML
2    Author:     Tobias Nipkow
3    Copyright   1995  TU Munich
4
5Generic case-splitter, suitable for most logics.
6Deals with equalities of the form ?P(f args) = ...
7where "f args" must be a first-order term without duplicate variables.
8*)
9
10signature SPLITTER_DATA =
11sig
12  val context       : Proof.context
13  val mk_eq         : thm -> thm
14  val meta_eq_to_iff: thm (* "x == y ==> x = y"                      *)
15  val iffD          : thm (* "[| P = Q; Q |] ==> P"                  *)
16  val disjE         : thm (* "[| P | Q; P ==> R; Q ==> R |] ==> R"   *)
17  val conjE         : thm (* "[| P & Q; [| P; Q |] ==> R |] ==> R"   *)
18  val exE           : thm (* "[| EX x. P x; !!x. P x ==> Q |] ==> Q" *)
19  val contrapos     : thm (* "[| ~ Q; P ==> Q |] ==> ~ P"            *)
20  val contrapos2    : thm (* "[| Q; ~ P ==> ~ Q |] ==> P"            *)
21  val notnotD       : thm (* "~ ~ P ==> P"                           *)
22  val safe_tac      : Proof.context -> tactic
23end
24
25signature SPLITTER =
26sig
27  (* somewhat more internal functions *)
28  val cmap_of_split_thms: thm list -> (string * (typ * term * thm * typ * int) list) list
29  val split_posns: (string * (typ * term * thm * typ * int) list) list ->
30    theory -> typ list -> term -> (thm * (typ * typ * int list) list * int list * typ * term) list
31    (* first argument is a "cmap", returns a list of "split packs" *)
32  (* the "real" interface, providing a number of tactics *)
33  val split_tac       : Proof.context -> thm list -> int -> tactic
34  val split_inside_tac: Proof.context -> thm list -> int -> tactic
35  val split_asm_tac   : Proof.context -> thm list -> int -> tactic
36  val add_split: thm -> Proof.context -> Proof.context
37  val add_split_bang: thm -> Proof.context -> Proof.context
38  val del_split: thm -> Proof.context -> Proof.context
39  val split_modifiers : Method.modifier parser list
40end;
41
42functor Splitter(Data: SPLITTER_DATA): SPLITTER =
43struct
44
45val Const (const_not, _) $ _ =
46  Object_Logic.drop_judgment Data.context
47    (#1 (Logic.dest_implies (Thm.prop_of Data.notnotD)));
48
49val Const (const_or , _) $ _ $ _ =
50  Object_Logic.drop_judgment Data.context
51    (#1 (Logic.dest_implies (Thm.prop_of Data.disjE)));
52
53val const_Trueprop = Object_Logic.judgment_name Data.context;
54
55
56fun split_format_err () = error "Wrong format for split rule";
57
58fun split_thm_info thm =
59  (case Thm.concl_of (Data.mk_eq thm) of
60    Const(@{const_name Pure.eq}, _) $ (Var _ $ t) $ c =>
61      (case strip_comb t of
62        (Const p, _) => (p, case c of (Const (s, _) $ _) => s = const_not | _ => false)
63      | _ => split_format_err ())
64  | _ => split_format_err ());
65
66fun cmap_of_split_thms thms =
67let
68  val splits = map Data.mk_eq thms
69  fun add_thm thm cmap =
70    (case Thm.concl_of thm of _ $ (t as _ $ lhs) $ _ =>
71       (case strip_comb lhs of (Const(a,aT),args) =>
72          let val info = (aT,lhs,thm,fastype_of t,length args)
73          in case AList.lookup (op =) cmap a of
74               SOME infos => AList.update (op =) (a, info::infos) cmap
75             | NONE => (a,[info])::cmap
76          end
77        | _ => split_format_err())
78     | _ => split_format_err())
79in
80  fold add_thm splits []
81end;
82
83val abss = fold (Term.abs o pair "");
84
85(* ------------------------------------------------------------------------- *)
86(* mk_case_split_tac                                                         *)
87(* ------------------------------------------------------------------------- *)
88
89fun mk_case_split_tac order =
90let
91
92(************************************************************
93   Create lift-theorem "trlift" :
94
95   [| !!x. Q x == R x; P(%x. R x) == C |] ==> P (%x. Q x) == C
96
97*************************************************************)
98
99val meta_iffD = Data.meta_eq_to_iff RS Data.iffD;  (* (P == Q) ==> Q ==> P *)
100
101val lift = Goal.prove_global @{theory Pure} ["P", "Q", "R"]
102  [Syntax.read_prop_global @{theory Pure} "!!x :: 'b. Q(x) == R(x) :: 'c"]
103  (Syntax.read_prop_global @{theory Pure} "P(%x. Q(x)) == P(%x. R(x))")
104  (fn {context = ctxt, prems} =>
105    rewrite_goals_tac ctxt prems THEN resolve_tac ctxt [reflexive_thm] 1)
106
107val _ $ _ $ (_ $ (_ $ abs_lift) $ _) = Thm.prop_of lift;
108
109val trlift = lift RS transitive_thm;
110
111
112(************************************************************************
113   Set up term for instantiation of P in the lift-theorem
114
115   t     : lefthand side of meta-equality in subgoal
116           the lift theorem is applied to (see select)
117   pos   : "path" leading to abstraction, coded as a list
118   T     : type of body of P(...)
119*************************************************************************)
120
121fun mk_cntxt t pos T =
122  let
123    fun down [] t = (Bound 0, t)
124      | down (p :: ps) t =
125          let
126            val (h, ts) = strip_comb t
127            val (ts1, u :: ts2) = chop p ts
128            val (u1, u2) = down ps u
129          in
130            (list_comb (incr_boundvars 1 h,
131               map (incr_boundvars 1) ts1 @ u1 ::
132               map (incr_boundvars 1) ts2),
133             u2)
134          end;
135    val (u1, u2) = down (rev pos) t
136  in (Abs ("", T, u1), u2) end;
137
138
139(************************************************************************
140   Set up term for instantiation of P in the split-theorem
141   P(...) == rhs
142
143   t     : lefthand side of meta-equality in subgoal
144           the split theorem is applied to (see select)
145   T     : type of body of P(...)
146   tt    : the term  Const(key,..) $ ...
147*************************************************************************)
148
149fun mk_cntxt_splitthm t tt T =
150  let fun repl lev t =
151    if Envir.aeconv(incr_boundvars lev tt, t) then Bound lev
152    else case t of
153        (Abs (v, T2, t)) => Abs (v, T2, repl (lev+1) t)
154      | (Bound i) => Bound (if i>=lev then i+1 else i)
155      | (t1 $ t2) => (repl lev t1) $ (repl lev t2)
156      | t => t
157  in Abs("", T, repl 0 t) end;
158
159
160(* add all loose bound variables in t to list is *)
161fun add_lbnos t is = add_loose_bnos (t, 0, is);
162
163(* check if the innermost abstraction that needs to be removed
164   has a body of type T; otherwise the expansion thm will fail later on
165*)
166fun type_test (T, lbnos, apsns) =
167  let val (_, U: typ, _) = nth apsns (foldl1 Int.min lbnos)
168  in T = U end;
169
170(*************************************************************************
171   Create a "split_pack".
172
173   thm   : the relevant split-theorem, i.e. P(...) == rhs , where P(...)
174           is of the form
175           P( Const(key,...) $ t_1 $ ... $ t_n )      (e.g. key = "if")
176   T     : type of P(...)
177   T'    : type of term to be scanned
178   n     : number of arguments expected by Const(key,...)
179   ts    : list of arguments actually found
180   apsns : list of tuples of the form (T,U,pos), one tuple for each
181           abstraction that is encountered on the way to the position where
182           Const(key, ...) $ ...  occurs, where
183           T   : type of the variable bound by the abstraction
184           U   : type of the abstraction's body
185           pos : "path" leading to the body of the abstraction
186   pos   : "path" leading to the position where Const(key, ...) $ ...  occurs.
187   TB    : type of  Const(key,...) $ t_1 $ ... $ t_n
188   t     : the term Const(key,...) $ t_1 $ ... $ t_n
189
190   A split pack is a tuple of the form
191   (thm, apsns, pos, TB, tt)
192   Note : apsns is reversed, so that the outermost quantifier's position
193          comes first ! If the terms in ts don't contain variables bound
194          by other than meta-quantifiers, apsns is empty, because no further
195          lifting is required before applying the split-theorem.
196******************************************************************************)
197
198fun mk_split_pack (thm, T: typ, T', n, ts, apsns, pos, TB, t) =
199  if n > length ts then []
200  else let val lev = length apsns
201           val lbnos = fold add_lbnos (take n ts) []
202           val flbnos = filter (fn i => i < lev) lbnos
203           val tt = incr_boundvars (~lev) t
204       in if null flbnos then
205            if T = T' then [(thm,[],pos,TB,tt)] else []
206          else if type_test(T,flbnos,apsns) then [(thm, rev apsns,pos,TB,tt)]
207               else []
208       end;
209
210
211(****************************************************************************
212   Recursively scans term for occurrences of Const(key,...) $ ...
213   Returns a list of "split-packs" (one for each occurrence of Const(key,...) )
214
215   cmap : association list of split-theorems that should be tried.
216          The elements have the format (key,(thm,T,n)) , where
217          key : the theorem's key constant ( Const(key,...) $ ... )
218          thm : the theorem itself
219          T   : type of P( Const(key,...) $ ... )
220          n   : number of arguments expected by Const(key,...)
221   Ts   : types of parameters
222   t    : the term to be scanned
223******************************************************************************)
224
225(* Simplified first-order matching;
226   assumes that all Vars in the pattern are distinct;
227   see Pure/pattern.ML for the full version;
228*)
229local
230  exception MATCH
231in
232  fun typ_match thy (tyenv, TU) = Sign.typ_match thy TU tyenv
233    handle Type.TYPE_MATCH => raise MATCH;
234
235  fun fomatch thy args =
236    let
237      fun mtch tyinsts = fn
238          (Ts, Var(_,T), t) =>
239            typ_match thy (tyinsts, (T, fastype_of1(Ts,t)))
240        | (_, Free (a,T), Free (b,U)) =>
241            if a=b then typ_match thy (tyinsts,(T,U)) else raise MATCH
242        | (_, Const (a,T), Const (b,U)) =>
243            if a=b then typ_match thy (tyinsts,(T,U)) else raise MATCH
244        | (_, Bound i, Bound j) =>
245            if i=j then tyinsts else raise MATCH
246        | (Ts, Abs(_,T,t), Abs(_,U,u)) =>
247            mtch (typ_match thy (tyinsts,(T,U))) (U::Ts,t,u)
248        | (Ts, f$t, g$u) =>
249            mtch (mtch tyinsts (Ts,f,g)) (Ts, t, u)
250        | _ => raise MATCH
251    in (mtch Vartab.empty args; true) handle MATCH => false end;
252end;
253
254fun split_posns (cmap : (string * (typ * term * thm * typ * int) list) list) thy Ts t =
255  let
256    val T' = fastype_of1 (Ts, t);
257    fun posns Ts pos apsns (Abs (_, T, t)) =
258          let val U = fastype_of1 (T::Ts,t)
259          in posns (T::Ts) (0::pos) ((T, U, pos)::apsns) t end
260      | posns Ts pos apsns t =
261          let
262            val (h, ts) = strip_comb t
263            fun iter t (i, a) = (i+1, (posns Ts (i::pos) apsns t) @ a);
264            val a =
265              case h of
266                Const(c, cT) =>
267                  let fun find [] = []
268                        | find ((gcT, pat, thm, T, n)::tups) =
269                            let val t2 = list_comb (h, take n ts) in
270                              if Sign.typ_instance thy (cT, gcT) andalso fomatch thy (Ts, pat, t2)
271                              then mk_split_pack(thm,T,T',n,ts,apsns,pos,type_of1(Ts,t2),t2)
272                              else find tups
273                            end
274                  in find (these (AList.lookup (op =) cmap c)) end
275              | _ => []
276          in snd (fold iter ts (0, a)) end
277  in posns Ts [] [] t end;
278
279fun shorter ((_,ps,pos,_,_), (_,qs,qos,_,_)) =
280  prod_ord (int_ord o apply2 length) (order o apply2 length)
281    ((ps, pos), (qs, qos));
282
283
284(************************************************************
285   call split_posns with appropriate parameters
286*************************************************************)
287
288fun select thy cmap state i =
289  let
290    val goal = Thm.term_of (Thm.cprem_of state i);
291    val Ts = rev (map #2 (Logic.strip_params goal));
292    val _ $ t $ _ = Logic.strip_assums_concl goal;
293  in (Ts, t, sort shorter (split_posns cmap thy Ts t)) end;
294
295fun exported_split_posns cmap thy Ts t =
296  sort shorter (split_posns cmap thy Ts t);
297
298(*************************************************************
299   instantiate lift theorem
300
301   if t is of the form
302   ... ( Const(...,...) $ Abs( .... ) ) ...
303   then
304   P = %a.  ... ( Const(...,...) $ a ) ...
305   where a has type T --> U
306
307   Ts      : types of parameters
308   t       : lefthand side of meta-equality in subgoal
309             the split theorem is applied to (see cmap)
310   T,U,pos : see mk_split_pack
311   state   : current proof state
312   i       : no. of subgoal
313**************************************************************)
314
315fun inst_lift ctxt Ts t (T, U, pos) state i =
316  let
317    val (cntxt, u) = mk_cntxt t pos (T --> U);
318    val trlift' = Thm.lift_rule (Thm.cprem_of state i)
319      (Thm.rename_boundvars abs_lift u trlift);
320    val (Var (P, _), _) =
321      strip_comb (fst (Logic.dest_equals
322        (Logic.strip_assums_concl (Thm.prop_of trlift'))));
323  in infer_instantiate ctxt [(P, Thm.cterm_of ctxt (abss Ts cntxt))] trlift' end;
324
325
326(*************************************************************
327   instantiate split theorem
328
329   Ts    : types of parameters
330   t     : lefthand side of meta-equality in subgoal
331           the split theorem is applied to (see cmap)
332   tt    : the term  Const(key,..) $ ...
333   thm   : the split theorem
334   TB    : type of body of P(...)
335   state : current proof state
336   i     : number of subgoal
337**************************************************************)
338
339fun inst_split ctxt Ts t tt thm TB state i =
340  let
341    val thm' = Thm.lift_rule (Thm.cprem_of state i) thm;
342    val (Var (P, _), _) =
343      strip_comb (fst (Logic.dest_equals
344        (Logic.strip_assums_concl (Thm.prop_of thm'))));
345    val cntxt = mk_cntxt_splitthm t tt TB;
346  in infer_instantiate ctxt [(P, Thm.cterm_of ctxt (abss Ts cntxt))] thm' end;
347
348
349(*****************************************************************************
350   The split-tactic
351
352   splits : list of split-theorems to be tried
353   i      : number of subgoal the tactic should be applied to
354*****************************************************************************)
355
356fun split_tac _ [] i = no_tac
357  | split_tac ctxt splits i =
358  let val cmap = cmap_of_split_thms splits
359      fun lift_tac Ts t p st = compose_tac ctxt (false, inst_lift ctxt Ts t p st i, 2) i st
360      fun lift_split_tac state =
361            let val (Ts, t, splits) = select (Proof_Context.theory_of ctxt) cmap state i
362            in case splits of
363                 [] => no_tac state
364               | (thm, apsns, pos, TB, tt)::_ =>
365                   (case apsns of
366                      [] =>
367                        compose_tac ctxt (false, inst_split ctxt Ts t tt thm TB state i, 0) i state
368                    | p::_ => EVERY [lift_tac Ts t p,
369                                     resolve_tac ctxt [reflexive_thm] (i+1),
370                                     lift_split_tac] state)
371            end
372  in COND (has_fewer_prems i) no_tac
373          (resolve_tac ctxt [meta_iffD] i THEN lift_split_tac)
374  end;
375
376in (split_tac, exported_split_posns) end;  (* mk_case_split_tac *)
377
378
379val (split_tac, split_posns) = mk_case_split_tac int_ord;
380
381val (split_inside_tac, _) = mk_case_split_tac (rev_order o int_ord);
382
383
384(*****************************************************************************
385   The split-tactic for premises
386
387   splits : list of split-theorems to be tried
388****************************************************************************)
389fun split_asm_tac _ [] = K no_tac
390  | split_asm_tac ctxt splits =
391
392  let val cname_list = map (fst o fst o split_thm_info) splits;
393      fun tac (t,i) =
394          let val n = find_index (exists_Const (member (op =) cname_list o #1))
395                                 (Logic.strip_assums_hyp t);
396              fun first_prem_is_disj (Const (@{const_name Pure.imp}, _) $ (Const (c, _)
397                    $ (Const (s, _) $ _ $ _ )) $ _ ) = c = const_Trueprop andalso s = const_or
398              |   first_prem_is_disj (Const(@{const_name Pure.all},_)$Abs(_,_,t)) =
399                                        first_prem_is_disj t
400              |   first_prem_is_disj _ = false;
401      (* does not work properly if the split variable is bound by a quantifier *)
402              fun flat_prems_tac i = SUBGOAL (fn (t,i) =>
403                           (if first_prem_is_disj t
404                            then EVERY[eresolve_tac ctxt [Data.disjE] i, rotate_tac ~1 i,
405                                       rotate_tac ~1  (i+1),
406                                       flat_prems_tac (i+1)]
407                            else all_tac)
408                           THEN REPEAT (eresolve_tac ctxt [Data.conjE,Data.exE] i)
409                           THEN REPEAT (dresolve_tac ctxt [Data.notnotD]   i)) i;
410          in if n<0 then  no_tac  else (DETERM (EVERY'
411                [rotate_tac n, eresolve_tac ctxt [Data.contrapos2],
412                 split_tac ctxt splits,
413                 rotate_tac ~1, eresolve_tac ctxt [Data.contrapos], rotate_tac ~1,
414                 flat_prems_tac] i))
415          end;
416  in SUBGOAL tac
417  end;
418
419fun gen_split_tac _ [] = K no_tac
420  | gen_split_tac ctxt (split::splits) =
421      let val (_,asm) = split_thm_info split
422      in (if asm then split_asm_tac else split_tac) ctxt [split] ORELSE'
423         gen_split_tac ctxt splits
424      end;
425
426
427(** declare split rules **)
428
429(* add_split / del_split *)
430
431fun string_of_typ (Type (s, Ts)) =
432      (if null Ts then "" else enclose "(" ")" (commas (map string_of_typ Ts))) ^ s
433  | string_of_typ _ = "_";
434
435fun split_name (name, T) asm = "split " ^
436  (if asm then "asm " else "") ^ name ^ " :: " ^ string_of_typ T;
437
438fun gen_add_split bang split ctxt =
439  let
440    val (name, asm) = split_thm_info split
441    val split0 = Thm.trim_context split
442    fun tac ctxt' =
443      let val split' = Thm.transfer' ctxt' split0 in
444        (if asm then split_asm_tac ctxt' [split']
445         else if bang
446              then split_tac ctxt' [split'] THEN_ALL_NEW
447                   TRY o (SELECT_GOAL (Data.safe_tac ctxt'))
448              else split_tac ctxt' [split'])
449      end
450  in Simplifier.addloop (ctxt, (split_name name asm, tac)) end;
451
452val add_split = gen_add_split false;
453val add_split_bang = gen_add_split true;
454
455fun del_split split ctxt =
456  let val (name, asm) = split_thm_info split
457  in Simplifier.delloop (ctxt, split_name name asm) end;
458
459
460(* attributes *)
461
462val splitN = "split";
463
464fun split_add bang = Simplifier.attrib (gen_add_split bang);
465val split_del = Simplifier.attrib del_split;
466
467val add_del = Scan.lift
468  (Args.bang >> K (split_add true)
469   || Args.del >> K split_del
470   || Scan.succeed (split_add false));
471
472val _ = Theory.setup
473  (Attrib.setup @{binding split} add_del "declare case split rule");
474
475
476(* methods *)
477
478val split_modifiers =
479 [Args.$$$ splitN -- Args.colon >> K (Method.modifier (split_add false) \<^here>),
480  Args.$$$ splitN -- Args.bang_colon >> K (Method.modifier (split_add true) \<^here>),
481  Args.$$$ splitN -- Args.del -- Args.colon >> K (Method.modifier split_del \<^here>)];
482
483val _ =
484  Theory.setup
485    (Method.setup @{binding split}
486      (Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (CHANGED_PROP o gen_split_tac ctxt ths)))
487      "apply case split rule");
488
489end;
490