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