TFL/casesplit.ML
changeset 15150 c7af682b9ee5
child 15250 217bececa2bd
equal deleted inserted replaced
15149:c5c4884634b7 15150:c7af682b9ee5
       
     1 (* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
       
     2 (*  Title:      TFL/casesplit.ML
       
     3     Author:     Lucas Dixon, University of Edinburgh
       
     4                 lucas.dixon@ed.ac.uk
       
     5     Date:       17 Aug 2004
       
     6 *)
       
     7 (* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
       
     8 (*  DESCRIPTION:
       
     9 
       
    10     A structure that defines a tactic to program case splits. 
       
    11 
       
    12     casesplit_free :
       
    13       string * Term.type -> int -> Thm.thm -> Thm.thm Seq.seq
       
    14 
       
    15     casesplit_name : 
       
    16       string -> int -> Thm.thm -> Thm.thm Seq.seq
       
    17 
       
    18     These use the induction theorem associated with the recursive data
       
    19     type to be split. 
       
    20 
       
    21     The structure includes a function to try and recursively split a
       
    22     conjecture into a list sub-theorems: 
       
    23 
       
    24     splitto : Thm.thm list -> Thm.thm -> Thm.thm
       
    25 *)
       
    26 (* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
       
    27 
       
    28 (* logic-specific *)
       
    29 signature CASE_SPLIT_DATA =
       
    30 sig
       
    31   val dest_Trueprop : Term.term -> Term.term
       
    32   val mk_Trueprop : Term.term -> Term.term
       
    33   val read_cterm : Sign.sg -> string -> Thm.cterm
       
    34 end;
       
    35 
       
    36 (* for HOL *)
       
    37 structure CaseSplitData_HOL : CASE_SPLIT_DATA = 
       
    38 struct
       
    39 val dest_Trueprop = HOLogic.dest_Trueprop;
       
    40 val mk_Trueprop = HOLogic.mk_Trueprop;
       
    41 val read_cterm = HOLogic.read_cterm;
       
    42 end;
       
    43 
       
    44 
       
    45 signature CASE_SPLIT =
       
    46 sig
       
    47   (* failure to find a free to split on *)
       
    48   exception find_split_exp of string
       
    49 
       
    50   (* getting a case split thm from the induction thm *)
       
    51   val case_thm_of_ty : Sign.sg -> Term.typ -> Thm.thm
       
    52   val cases_thm_of_induct_thm : Thm.thm -> Thm.thm
       
    53 
       
    54   (* case split tactics *)
       
    55   val casesplit_free :
       
    56       string * Term.typ -> int -> Thm.thm -> Thm.thm Seq.seq
       
    57   val casesplit_name : string -> int -> Thm.thm -> Thm.thm Seq.seq
       
    58 
       
    59   (* finding a free var to split *)
       
    60   val find_term_split :
       
    61       Term.term * Term.term -> (string * Term.typ) Library.option
       
    62   val find_thm_split :
       
    63       Thm.thm -> int -> Thm.thm -> (string * Term.typ) Library.option
       
    64   val find_thms_split :
       
    65       Thm.thm list -> int -> Thm.thm -> (string * Term.typ) Library.option
       
    66 
       
    67   (* try to recursively split conjectured thm to given list of thms *)
       
    68   val splitto : Thm.thm list -> Thm.thm -> Thm.thm
       
    69 
       
    70   (* for use with the recdef package *)
       
    71   val derive_init_eqs :
       
    72       Sign.sg ->
       
    73       (Thm.thm * int) list -> Term.term list -> (Thm.thm * int) list
       
    74 end;
       
    75 
       
    76 functor CaseSplitFUN(Data : CASE_SPLIT_DATA) =
       
    77 struct
       
    78 
       
    79 (* beta-eta contract the theorem *)
       
    80 fun beta_eta_contract thm = 
       
    81     let
       
    82       val thm2 = equal_elim (Thm.beta_conversion true (Thm.cprop_of thm)) thm
       
    83       val thm3 = equal_elim (Thm.eta_conversion (Thm.cprop_of thm2)) thm2
       
    84     in thm3 end;
       
    85 
       
    86 (* make a casethm from an induction thm *)
       
    87 val cases_thm_of_induct_thm = 
       
    88      Seq.hd o (ALLGOALS (fn i => REPEAT (etac Drule.thin_rl i)));
       
    89 
       
    90 (* get the case_thm (my version) from a type *)
       
    91 fun case_thm_of_ty sgn ty  = 
       
    92     let 
       
    93       val dtypestab = DatatypePackage.get_datatypes_sg sgn;
       
    94       val ty_str = case ty of 
       
    95                      Type(ty_str, _) => ty_str
       
    96                    | TFree(s,_)  => raise ERROR_MESSAGE 
       
    97                                             ("Free type: " ^ s)   
       
    98                    | TVar((s,i),_) => raise ERROR_MESSAGE 
       
    99                                             ("Free variable: " ^ s)   
       
   100       val dt = case (Symtab.lookup (dtypestab,ty_str))
       
   101                 of Some dt => dt
       
   102                  | None => raise ERROR_MESSAGE ("Not a Datatype: " ^ ty_str)
       
   103     in
       
   104       cases_thm_of_induct_thm (#induction dt)
       
   105     end;
       
   106 
       
   107 (* 
       
   108  val ty = (snd o hd o map Term.dest_Free o Term.term_frees) t;  
       
   109 *)
       
   110 
       
   111 
       
   112 (* for use when there are no prems to the subgoal *)
       
   113 (* does a case split on the given variable *)
       
   114 fun mk_casesplit_goal_thm sgn (vstr,ty) gt = 
       
   115     let 
       
   116       val x = Free(vstr,ty)
       
   117       val abst = Abs(vstr, ty, Term.abstract_over (x, gt));
       
   118 
       
   119       val ctermify = Thm.cterm_of sgn;
       
   120       val ctypify = Thm.ctyp_of sgn;
       
   121       val case_thm = case_thm_of_ty sgn ty;
       
   122 
       
   123       val abs_ct = ctermify abst;
       
   124       val free_ct = ctermify x;
       
   125 
       
   126       val casethm_vars = rev (Term.term_vars (Thm.concl_of case_thm));
       
   127        
       
   128       val tsig = Sign.tsig_of sgn;
       
   129       val casethm_tvars = Term.term_tvars (Thm.concl_of case_thm);
       
   130       val (Pv, Dv, type_insts) = 
       
   131           case (Thm.concl_of case_thm) of 
       
   132             (_ $ ((Pv as Var(P,Pty)) $ (Dv as Var(D, Dty)))) => 
       
   133             (Pv, Dv, 
       
   134              Vartab.dest (Type.typ_match tsig (Vartab.empty, (Dty, ty))))
       
   135           | _ => raise ERROR_MESSAGE ("not a valid case thm");
       
   136       val type_cinsts = map (apsnd ctypify) type_insts;
       
   137       val cPv = ctermify (Sign.inst_term_tvars sgn type_insts Pv);
       
   138       val cDv = ctermify (Sign.inst_term_tvars sgn type_insts Dv);
       
   139     in
       
   140       (beta_eta_contract 
       
   141          (case_thm
       
   142             |> Thm.instantiate (type_cinsts, []) 
       
   143             |> Thm.instantiate ([], [(cPv, abs_ct), (cDv, free_ct)])))
       
   144     end;
       
   145 
       
   146 
       
   147 (* for use when there are no prems to the subgoal *)
       
   148 (* does a case split on the given variable (Free fv) *)
       
   149 fun casesplit_free fv i th = 
       
   150     let 
       
   151       val gt = Data.dest_Trueprop (nth_elem( i - 1, Thm.prems_of th));
       
   152       val sgn = Thm.sign_of_thm th;
       
   153     in 
       
   154       Tactic.rtac (mk_casesplit_goal_thm sgn fv gt) i th
       
   155     end;
       
   156 
       
   157 (* for use when there are no prems to the subgoal *)
       
   158 (* does a case split on the given variable *)
       
   159 fun casesplit_name vstr i th = 
       
   160     let 
       
   161       val gt = Data.dest_Trueprop (nth_elem( i - 1, Thm.prems_of th));
       
   162       val freets = Term.term_frees gt;
       
   163       fun getter x = let val (n,ty) = Term.dest_Free x in 
       
   164                        if vstr = n then Some (n,ty) else None end;
       
   165       val (n,ty) = case Library.get_first getter freets 
       
   166                 of Some (n, ty) => (n, ty)
       
   167                  | _ => raise ERROR_MESSAGE ("no such variable " ^ vstr);
       
   168       val sgn = Thm.sign_of_thm th;
       
   169     in 
       
   170       Tactic.rtac (mk_casesplit_goal_thm sgn (n,ty) gt) i th
       
   171     end;
       
   172 
       
   173 
       
   174 (* small example: 
       
   175 Goal "P (x :: nat) & (C y --> Q (y :: nat))";
       
   176 by (rtac (thm "conjI") 1);
       
   177 val th = topthm();
       
   178 val i = 2;
       
   179 val vstr = "y";
       
   180 
       
   181 by (casesplit_name "y" 2);
       
   182 
       
   183 val th = topthm();
       
   184 val i = 1;
       
   185 val th' = casesplit_name "x" i th;
       
   186 *)
       
   187 
       
   188 
       
   189 (* the find_XXX_split functions are simply doing a lightwieght (I
       
   190 think) term matching equivalent to find where to do the next split *)
       
   191 
       
   192 (* assuming two twems are identical except for a free in one at a
       
   193 subterm, or constant in another, ie assume that one term is a plit of
       
   194 another, then gives back the free variable that has been split. *)
       
   195 exception find_split_exp of string
       
   196 fun find_term_split (Free v, _ $ _) = Some v
       
   197   | find_term_split (Free v, Const _) = Some v
       
   198   | find_term_split (Free v, Abs _) = Some v (* do we really want this case? *)
       
   199   | find_term_split (a $ b, a2 $ b2) = 
       
   200     (case find_term_split (a, a2) of 
       
   201        None => find_term_split (b,b2)  
       
   202      | vopt => vopt)
       
   203   | find_term_split (Abs(_,ty,t1), Abs(_,ty2,t2)) = 
       
   204     find_term_split (t1, t2)
       
   205   | find_term_split (Const (x,ty), Const(x2,ty2)) = 
       
   206     if x = x2 then None else (* keep searching *)
       
   207     raise find_split_exp (* stop now *)
       
   208             "Terms are not identical upto a free varaible! (Consts)"
       
   209   | find_term_split (Bound i, Bound j) =     
       
   210     if i = j then None else (* keep searching *)
       
   211     raise find_split_exp (* stop now *)
       
   212             "Terms are not identical upto a free varaible! (Bound)"
       
   213   | find_term_split (a, b) = 
       
   214     raise find_split_exp (* stop now *)
       
   215             "Terms are not identical upto a free varaible! (Other)";
       
   216 
       
   217 (* assume that "splitth" is a case split form of subgoal i of "genth",
       
   218 then look for a free variable to split, breaking the subgoal closer to
       
   219 splitth. *)
       
   220 fun find_thm_split splitth i genth =
       
   221     find_term_split (Logic.get_goal (Thm.prop_of genth) i, 
       
   222                      Thm.concl_of splitth) handle find_split_exp _ => None;
       
   223 
       
   224 (* as above but searches "splitths" for a theorem that suggest a case split *)
       
   225 fun find_thms_split splitths i genth =
       
   226     Library.get_first (fn sth => find_thm_split sth i genth) splitths;
       
   227 
       
   228 
       
   229 (* split the subgoal i of "genth" until we get to a member of
       
   230 splitths. Assumes that genth will be a general form of splitths, that
       
   231 can be case-split, as needed. Otherwise fails. Note: We assume that
       
   232 all of "splitths" are aplit to the same level, and thus it doesn't
       
   233 matter which one we choose to look for the next split. Simply add
       
   234 search on splitthms and plit variable, to change this.  *)
       
   235 (* Note: possible efficiency measure: when a case theorem is no longer
       
   236 useful, drop it? *)
       
   237 (* Note: This should not be a separate tactic but integrated into the
       
   238 case split done during recdef's case analysis, this would avoid us
       
   239 having to (re)search for variables to split. *)
       
   240 fun splitto splitths genth = 
       
   241     let 
       
   242       val _ = assert (not (null splitths)) "splitto: no given splitths";
       
   243       val sgn = Thm.sign_of_thm genth;
       
   244 
       
   245       (* check if we are a member of splitths - FIXME: quicker and 
       
   246       more flexible with discrim net. *)
       
   247       fun solve_by_splitth th split = biresolution false [(false,split)] 1 th;
       
   248 
       
   249       fun split th = 
       
   250           (case find_thms_split splitths 1 th of 
       
   251              None => raise ERROR_MESSAGE "splitto: cannot find variable to split on"
       
   252             | Some v => 
       
   253              let 
       
   254                val gt = Data.dest_Trueprop (nth_elem(0, Thm.prems_of th));
       
   255                val split_thm = mk_casesplit_goal_thm sgn v gt;
       
   256                val (subthms, expf) = IsaND.fixed_subgoal_thms split_thm;
       
   257              in 
       
   258                expf (map recsplitf subthms)
       
   259              end)
       
   260 
       
   261       and recsplitf th = 
       
   262           (* note: multiple unifiers! we only take the first element,
       
   263              probably fine -- there is probably only one anyway. *)
       
   264           (case Library.get_first (Seq.pull o solve_by_splitth th) splitths of
       
   265              None => split th
       
   266            | Some (solved_th, more) => solved_th)
       
   267     in
       
   268       recsplitf genth
       
   269     end;
       
   270 
       
   271 
       
   272 (* Note: We dont do this if wf conditions fail to be solved, as each
       
   273 case may have a different wf condition - we could group the conditions
       
   274 togeather and say that they must be true to solve the general case,
       
   275 but that would hide from the user which sub-case they were related
       
   276 to. Probably this is not important, and it would work fine, but I
       
   277 prefer leaving more fine grain control to the user. *)
       
   278 
       
   279 (* derive eqs, assuming strict, ie the rules have no assumptions = all
       
   280    the well-foundness conditions have been solved. *)
       
   281 local
       
   282   fun get_related_thms i = 
       
   283       mapfilter ((fn (r,x) => if x = i then Some r else None));
       
   284       
       
   285   fun solve_eq (th, [], i) = 
       
   286       raise ERROR_MESSAGE "derive_init_eqs: missing rules"
       
   287     | solve_eq (th, [a], i) = (a, i)
       
   288     | solve_eq (th, splitths as (_ :: _), i) = (splitto splitths th,i);
       
   289 in
       
   290 fun derive_init_eqs sgn rules eqs = 
       
   291     let 
       
   292       val eqths = map (Thm.trivial o (Thm.cterm_of sgn) o Data.mk_Trueprop) 
       
   293                       eqs
       
   294     in
       
   295       (rev o map solve_eq)
       
   296         (Library.foldln 
       
   297            (fn (e,i) => 
       
   298                (curry (op ::)) (e, (get_related_thms (i - 1) rules), i - 1)) 
       
   299            eqths [])
       
   300     end;
       
   301 end;
       
   302 (* 
       
   303     val (rs_hwfc, unhidefs) = Library.split_list (map hide_prems rules)
       
   304     (map2 (op |>) (ths, expfs))
       
   305 *)
       
   306 
       
   307 end;
       
   308 
       
   309 
       
   310 structure CaseSplit = CaseSplitFUN(CaseSplitData_HOL);