wenzelm@16978: (* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) paulson@15150: (* Title: TFL/casesplit.ML paulson@15150: Author: Lucas Dixon, University of Edinburgh paulson@15150: lucas.dixon@ed.ac.uk paulson@15150: Date: 17 Aug 2004 paulson@15150: *) wenzelm@16978: (* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) paulson@15150: (* DESCRIPTION: paulson@15150: wenzelm@16978: A structure that defines a tactic to program case splits. paulson@15150: paulson@15150: casesplit_free : wenzelm@16978: string * typ -> int -> thm -> thm Seq.seq paulson@15150: wenzelm@16978: casesplit_name : wenzelm@16978: string -> int -> thm -> thm Seq.seq paulson@15150: paulson@15150: These use the induction theorem associated with the recursive data wenzelm@16978: type to be split. paulson@15150: paulson@15150: The structure includes a function to try and recursively split a wenzelm@16978: conjecture into a list sub-theorems: paulson@15150: wenzelm@16978: splitto : thm list -> thm -> thm paulson@15150: *) wenzelm@16978: (* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) paulson@15150: paulson@15150: (* logic-specific *) paulson@15150: signature CASE_SPLIT_DATA = paulson@15150: sig wenzelm@16978: val dest_Trueprop : term -> term wenzelm@16978: val mk_Trueprop : term -> term wenzelm@16978: val atomize : thm list wenzelm@18479: val rulify : thm list paulson@15150: end; paulson@15150: wenzelm@16978: structure CaseSplitData_HOL : CASE_SPLIT_DATA = paulson@15150: struct paulson@15150: val dest_Trueprop = HOLogic.dest_Trueprop; paulson@15150: val mk_Trueprop = HOLogic.mk_Trueprop; dixon@15250: dixon@15250: val atomize = thms "induct_atomize"; wenzelm@18479: val rulify = thms "induct_rulify"; wenzelm@18479: val rulify_fallback = thms "induct_rulify_fallback"; dixon@15250: paulson@15150: end; paulson@15150: paulson@15150: paulson@15150: signature CASE_SPLIT = paulson@15150: sig paulson@15150: (* failure to find a free to split on *) paulson@15150: exception find_split_exp of string paulson@15150: paulson@15150: (* getting a case split thm from the induction thm *) wenzelm@16978: val case_thm_of_ty : theory -> typ -> thm wenzelm@16978: val cases_thm_of_induct_thm : thm -> thm paulson@15150: paulson@15150: (* case split tactics *) paulson@15150: val casesplit_free : wenzelm@16978: string * typ -> int -> thm -> thm Seq.seq wenzelm@16978: val casesplit_name : string -> int -> thm -> thm Seq.seq paulson@15150: paulson@15150: (* finding a free var to split *) paulson@15150: val find_term_split : wenzelm@16978: term * term -> (string * typ) option paulson@15150: val find_thm_split : wenzelm@16978: thm -> int -> thm -> (string * typ) option paulson@15150: val find_thms_split : wenzelm@16978: thm list -> int -> thm -> (string * typ) option paulson@15150: paulson@15150: (* try to recursively split conjectured thm to given list of thms *) wenzelm@16978: val splitto : thm list -> thm -> thm paulson@15150: paulson@15150: (* for use with the recdef package *) paulson@15150: val derive_init_eqs : wenzelm@16978: theory -> wenzelm@16978: (thm * int) list -> term list -> (thm * int) list paulson@15150: end; paulson@15150: paulson@15150: functor CaseSplitFUN(Data : CASE_SPLIT_DATA) = paulson@15150: struct paulson@15150: wenzelm@18479: val rulify_goals = Tactic.rewrite_goals_rule Data.rulify; dixon@15250: val atomize_goals = Tactic.rewrite_goals_rule Data.atomize; dixon@15250: paulson@15150: (* beta-eta contract the theorem *) wenzelm@16978: fun beta_eta_contract thm = paulson@15150: let paulson@15150: val thm2 = equal_elim (Thm.beta_conversion true (Thm.cprop_of thm)) thm paulson@15150: val thm3 = equal_elim (Thm.eta_conversion (Thm.cprop_of thm2)) thm2 paulson@15150: in thm3 end; paulson@15150: paulson@15150: (* make a casethm from an induction thm *) wenzelm@16978: val cases_thm_of_induct_thm = paulson@15150: Seq.hd o (ALLGOALS (fn i => REPEAT (etac Drule.thin_rl i))); paulson@15150: paulson@15150: (* get the case_thm (my version) from a type *) wenzelm@16978: fun case_thm_of_ty sgn ty = wenzelm@16978: let wenzelm@16425: val dtypestab = DatatypePackage.get_datatypes sgn; wenzelm@16978: val ty_str = case ty of paulson@15150: Type(ty_str, _) => ty_str wenzelm@18678: | TFree(s,_) => error ("Free type: " ^ s) wenzelm@18678: | TVar((s,i),_) => error ("Free variable: " ^ s) wenzelm@17412: val dt = case Symtab.lookup dtypestab ty_str skalberg@15531: of SOME dt => dt wenzelm@18678: | NONE => error ("Not a Datatype: " ^ ty_str) paulson@15150: in paulson@15150: cases_thm_of_induct_thm (#induction dt) paulson@15150: end; paulson@15150: wenzelm@16978: (* wenzelm@16978: val ty = (snd o hd o map Term.dest_Free o Term.term_frees) t; paulson@15150: *) paulson@15150: paulson@15150: paulson@15150: (* for use when there are no prems to the subgoal *) paulson@15150: (* does a case split on the given variable *) wenzelm@16978: fun mk_casesplit_goal_thm sgn (vstr,ty) gt = wenzelm@16978: let paulson@15150: val x = Free(vstr,ty) paulson@15150: val abst = Abs(vstr, ty, Term.abstract_over (x, gt)); paulson@15150: paulson@15150: val ctermify = Thm.cterm_of sgn; paulson@15150: val ctypify = Thm.ctyp_of sgn; paulson@15150: val case_thm = case_thm_of_ty sgn ty; paulson@15150: paulson@15150: val abs_ct = ctermify abst; paulson@15150: val free_ct = ctermify x; paulson@15150: paulson@15150: val casethm_vars = rev (Term.term_vars (Thm.concl_of case_thm)); wenzelm@16978: paulson@15150: val casethm_tvars = Term.term_tvars (Thm.concl_of case_thm); wenzelm@16978: val (Pv, Dv, type_insts) = wenzelm@16978: case (Thm.concl_of case_thm) of wenzelm@16978: (_ $ ((Pv as Var(P,Pty)) $ (Dv as Var(D, Dty)))) => wenzelm@16978: (Pv, Dv, wenzelm@16935: Sign.typ_match sgn (Dty, ty) Vartab.empty) wenzelm@18678: | _ => error "not a valid case thm"; berghofe@15798: val type_cinsts = map (fn (ixn, (S, T)) => (ctypify (TVar (ixn, S)), ctypify T)) berghofe@15798: (Vartab.dest type_insts); berghofe@15798: val cPv = ctermify (Envir.subst_TVars type_insts Pv); berghofe@15798: val cDv = ctermify (Envir.subst_TVars type_insts Dv); paulson@15150: in wenzelm@16978: (beta_eta_contract paulson@15150: (case_thm wenzelm@16978: |> Thm.instantiate (type_cinsts, []) paulson@15150: |> Thm.instantiate ([], [(cPv, abs_ct), (cDv, free_ct)]))) paulson@15150: end; paulson@15150: paulson@15150: paulson@15150: (* for use when there are no prems to the subgoal *) paulson@15150: (* does a case split on the given variable (Free fv) *) wenzelm@16978: fun casesplit_free fv i th = wenzelm@16978: let dixon@15250: val (subgoalth, exp) = IsaND.fix_alls i th; dixon@15250: val subgoalth' = atomize_goals subgoalth; dixon@15250: val gt = Data.dest_Trueprop (Logic.get_goal (Thm.prop_of subgoalth') 1); paulson@15150: val sgn = Thm.sign_of_thm th; dixon@15250: dixon@15250: val splitter_thm = mk_casesplit_goal_thm sgn fv gt; dixon@15250: val nsplits = Thm.nprems_of splitter_thm; dixon@15250: dixon@15250: val split_goal_th = splitter_thm RS subgoalth'; dixon@15250: val rulified_split_goal_th = rulify_goals split_goal_th; wenzelm@16978: in dixon@15250: IsaND.export_back exp rulified_split_goal_th paulson@15150: end; paulson@15150: dixon@15250: paulson@15150: (* for use when there are no prems to the subgoal *) paulson@15150: (* does a case split on the given variable *) wenzelm@16978: fun casesplit_name vstr i th = wenzelm@16978: let dixon@15250: val (subgoalth, exp) = IsaND.fix_alls i th; dixon@15250: val subgoalth' = atomize_goals subgoalth; dixon@15250: val gt = Data.dest_Trueprop (Logic.get_goal (Thm.prop_of subgoalth') 1); dixon@15250: paulson@15150: val freets = Term.term_frees gt; wenzelm@16978: fun getter x = wenzelm@16978: let val (n,ty) = Term.dest_Free x in wenzelm@20081: (if vstr = n orelse vstr = Name.dest_skolem n skalberg@15531: then SOME (n,ty) else NONE ) skalberg@15570: handle Fail _ => NONE (* dest_skolem *) dixon@15250: end; wenzelm@16978: val (n,ty) = case Library.get_first getter freets skalberg@15531: of SOME (n, ty) => (n, ty) wenzelm@18678: | _ => error ("no such variable " ^ vstr); paulson@15150: val sgn = Thm.sign_of_thm th; dixon@15250: dixon@15250: val splitter_thm = mk_casesplit_goal_thm sgn (n,ty) gt; dixon@15250: val nsplits = Thm.nprems_of splitter_thm; dixon@15250: dixon@15250: val split_goal_th = splitter_thm RS subgoalth'; dixon@15250: dixon@15250: val rulified_split_goal_th = rulify_goals split_goal_th; wenzelm@16978: in dixon@15250: IsaND.export_back exp rulified_split_goal_th paulson@15150: end; paulson@15150: paulson@15150: wenzelm@16978: (* small example: paulson@15150: Goal "P (x :: nat) & (C y --> Q (y :: nat))"; paulson@15150: by (rtac (thm "conjI") 1); paulson@15150: val th = topthm(); paulson@15150: val i = 2; paulson@15150: val vstr = "y"; paulson@15150: paulson@15150: by (casesplit_name "y" 2); paulson@15150: paulson@15150: val th = topthm(); paulson@15150: val i = 1; paulson@15150: val th' = casesplit_name "x" i th; paulson@15150: *) paulson@15150: paulson@15150: paulson@15150: (* the find_XXX_split functions are simply doing a lightwieght (I paulson@15150: think) term matching equivalent to find where to do the next split *) paulson@15150: paulson@15150: (* assuming two twems are identical except for a free in one at a paulson@15150: subterm, or constant in another, ie assume that one term is a plit of paulson@15150: another, then gives back the free variable that has been split. *) paulson@15150: exception find_split_exp of string skalberg@15531: fun find_term_split (Free v, _ $ _) = SOME v skalberg@15531: | find_term_split (Free v, Const _) = SOME v skalberg@15531: | find_term_split (Free v, Abs _) = SOME v (* do we really want this case? *) skalberg@15531: | find_term_split (Free v, Var _) = NONE (* keep searching *) wenzelm@16978: | find_term_split (a $ b, a2 $ b2) = wenzelm@16978: (case find_term_split (a, a2) of wenzelm@16978: NONE => find_term_split (b,b2) paulson@15150: | vopt => vopt) wenzelm@16978: | find_term_split (Abs(_,ty,t1), Abs(_,ty2,t2)) = paulson@15150: find_term_split (t1, t2) wenzelm@16978: | find_term_split (Const (x,ty), Const(x2,ty2)) = skalberg@15531: if x = x2 then NONE else (* keep searching *) paulson@15150: raise find_split_exp (* stop now *) paulson@15150: "Terms are not identical upto a free varaible! (Consts)" wenzelm@16978: | find_term_split (Bound i, Bound j) = skalberg@15531: if i = j then NONE else (* keep searching *) paulson@15150: raise find_split_exp (* stop now *) paulson@15150: "Terms are not identical upto a free varaible! (Bound)" wenzelm@16978: | find_term_split (a, b) = paulson@15150: raise find_split_exp (* stop now *) paulson@15150: "Terms are not identical upto a free varaible! (Other)"; paulson@15150: paulson@15150: (* assume that "splitth" is a case split form of subgoal i of "genth", paulson@15150: then look for a free variable to split, breaking the subgoal closer to paulson@15150: splitth. *) paulson@15150: fun find_thm_split splitth i genth = wenzelm@16978: find_term_split (Logic.get_goal (Thm.prop_of genth) i, skalberg@15531: Thm.concl_of splitth) handle find_split_exp _ => NONE; paulson@15150: paulson@15150: (* as above but searches "splitths" for a theorem that suggest a case split *) paulson@15150: fun find_thms_split splitths i genth = paulson@15150: Library.get_first (fn sth => find_thm_split sth i genth) splitths; paulson@15150: paulson@15150: paulson@15150: (* split the subgoal i of "genth" until we get to a member of paulson@15150: splitths. Assumes that genth will be a general form of splitths, that paulson@15150: can be case-split, as needed. Otherwise fails. Note: We assume that dixon@15250: all of "splitths" are split to the same level, and thus it doesn't paulson@15150: matter which one we choose to look for the next split. Simply add dixon@15250: search on splitthms and split variable, to change this. *) paulson@15150: (* Note: possible efficiency measure: when a case theorem is no longer paulson@15150: useful, drop it? *) paulson@15150: (* Note: This should not be a separate tactic but integrated into the paulson@15150: case split done during recdef's case analysis, this would avoid us paulson@15150: having to (re)search for variables to split. *) wenzelm@16978: fun splitto splitths genth = wenzelm@16978: let paulson@15150: val _ = assert (not (null splitths)) "splitto: no given splitths"; paulson@15150: val sgn = Thm.sign_of_thm genth; paulson@15150: wenzelm@16978: (* check if we are a member of splitths - FIXME: quicker and paulson@15150: more flexible with discrim net. *) wenzelm@16978: fun solve_by_splitth th split = dixon@15250: Thm.biresolution false [(false,split)] 1 th; paulson@15150: wenzelm@16978: fun split th = wenzelm@16978: (case find_thms_split splitths 1 th of wenzelm@16978: NONE => dixon@15250: (writeln "th:"; berghofe@15252: Display.print_thm th; writeln "split ths:"; berghofe@15252: Display.print_thms splitths; writeln "\n--"; wenzelm@18678: error "splitto: cannot find variable to split on") wenzelm@16978: | SOME v => wenzelm@16978: let skalberg@15570: val gt = Data.dest_Trueprop (List.nth(Thm.prems_of th, 0)); paulson@15150: val split_thm = mk_casesplit_goal_thm sgn v gt; paulson@15150: val (subthms, expf) = IsaND.fixed_subgoal_thms split_thm; wenzelm@16978: in paulson@15150: expf (map recsplitf subthms) paulson@15150: end) paulson@15150: wenzelm@16978: and recsplitf th = paulson@15150: (* note: multiple unifiers! we only take the first element, paulson@15150: probably fine -- there is probably only one anyway. *) paulson@15150: (case Library.get_first (Seq.pull o solve_by_splitth th) splitths of skalberg@15531: NONE => split th skalberg@15531: | SOME (solved_th, more) => solved_th) paulson@15150: in paulson@15150: recsplitf genth paulson@15150: end; paulson@15150: paulson@15150: paulson@15150: (* Note: We dont do this if wf conditions fail to be solved, as each paulson@15150: case may have a different wf condition - we could group the conditions paulson@15150: togeather and say that they must be true to solve the general case, paulson@15150: but that would hide from the user which sub-case they were related paulson@15150: to. Probably this is not important, and it would work fine, but I paulson@15150: prefer leaving more fine grain control to the user. *) paulson@15150: paulson@15150: (* derive eqs, assuming strict, ie the rules have no assumptions = all paulson@15150: the well-foundness conditions have been solved. *) wenzelm@16978: fun derive_init_eqs sgn rules eqs = haftmann@18050: let haftmann@18050: fun get_related_thms i = haftmann@18050: List.mapPartial ((fn (r, x) => if x = i then SOME r else NONE)); haftmann@18050: fun add_eq (i, e) xs = haftmann@18050: (e, (get_related_thms i rules), i) :: xs haftmann@18050: fun solve_eq (th, [], i) = wenzelm@18678: error "derive_init_eqs: missing rules" haftmann@18050: | solve_eq (th, [a], i) = (a, i) haftmann@18050: | solve_eq (th, splitths as (_ :: _), i) = (splitto splitths th, i); haftmann@18050: val eqths = haftmann@18050: map (Thm.trivial o Thm.cterm_of sgn o Data.mk_Trueprop) eqs; haftmann@18050: in haftmann@18050: [] haftmann@18050: |> fold_index add_eq eqths haftmann@18050: |> map solve_eq haftmann@18050: |> rev haftmann@18050: end; paulson@15150: paulson@15150: end; paulson@15150: paulson@15150: paulson@15150: structure CaseSplit = CaseSplitFUN(CaseSplitData_HOL);