| 60520 |      1 | (*  Title:      HOL/Tools/old_recdef.ML
 | 
|  |      2 |     Author:     Konrad Slind, Cambridge University Computer Laboratory
 | 
|  |      3 |     Author:     Lucas Dixon, University of Edinburgh
 | 
|  |      4 | 
 | 
|  |      5 | Old TFL/recdef package.
 | 
|  |      6 | *)
 | 
|  |      7 | 
 | 
|  |      8 | signature CASE_SPLIT =
 | 
|  |      9 | sig
 | 
|  |     10 |   (* try to recursively split conjectured thm to given list of thms *)
 | 
|  |     11 |   val splitto : Proof.context -> thm list -> thm -> thm
 | 
|  |     12 | end;
 | 
|  |     13 | 
 | 
|  |     14 | signature UTILS =
 | 
|  |     15 | sig
 | 
|  |     16 |   exception ERR of {module: string, func: string, mesg: string}
 | 
|  |     17 |   val end_itlist: ('a -> 'a -> 'a) -> 'a list -> 'a
 | 
|  |     18 |   val itlist2: ('a -> 'b -> 'c -> 'c) -> 'a list -> 'b list -> 'c -> 'c
 | 
|  |     19 |   val pluck: ('a -> bool) -> 'a list -> 'a * 'a list
 | 
|  |     20 |   val zip3: 'a list -> 'b list -> 'c list -> ('a*'b*'c) list
 | 
|  |     21 |   val take: ('a -> 'b) -> int * 'a list -> 'b list
 | 
|  |     22 | end;
 | 
|  |     23 | 
 | 
|  |     24 | signature USYNTAX =
 | 
|  |     25 | sig
 | 
|  |     26 |   datatype lambda = VAR   of {Name : string, Ty : typ}
 | 
|  |     27 |                   | CONST of {Name : string, Ty : typ}
 | 
|  |     28 |                   | COMB  of {Rator: term, Rand : term}
 | 
|  |     29 |                   | LAMB  of {Bvar : term, Body : term}
 | 
|  |     30 | 
 | 
|  |     31 |   val alpha : typ
 | 
|  |     32 | 
 | 
|  |     33 |   (* Types *)
 | 
|  |     34 |   val type_vars  : typ -> typ list
 | 
|  |     35 |   val type_varsl : typ list -> typ list
 | 
|  |     36 |   val mk_vartype : string -> typ
 | 
|  |     37 |   val is_vartype : typ -> bool
 | 
|  |     38 |   val strip_prod_type : typ -> typ list
 | 
|  |     39 | 
 | 
|  |     40 |   (* Terms *)
 | 
|  |     41 |   val free_vars_lr : term -> term list
 | 
|  |     42 |   val type_vars_in_term : term -> typ list
 | 
|  |     43 |   val dest_term  : term -> lambda
 | 
|  |     44 | 
 | 
|  |     45 |   (* Prelogic *)
 | 
|  |     46 |   val inst      : (typ*typ) list -> term -> term
 | 
|  |     47 | 
 | 
|  |     48 |   (* Construction routines *)
 | 
|  |     49 |   val mk_abs    :{Bvar  : term, Body : term} -> term
 | 
|  |     50 | 
 | 
|  |     51 |   val mk_imp    :{ant : term, conseq :  term} -> term
 | 
|  |     52 |   val mk_select :{Bvar : term, Body : term} -> term
 | 
|  |     53 |   val mk_forall :{Bvar : term, Body : term} -> term
 | 
|  |     54 |   val mk_exists :{Bvar : term, Body : term} -> term
 | 
|  |     55 |   val mk_conj   :{conj1 : term, conj2 : term} -> term
 | 
|  |     56 |   val mk_disj   :{disj1 : term, disj2 : term} -> term
 | 
|  |     57 |   val mk_pabs   :{varstruct : term, body : term} -> term
 | 
|  |     58 | 
 | 
|  |     59 |   (* Destruction routines *)
 | 
|  |     60 |   val dest_const: term -> {Name : string, Ty : typ}
 | 
|  |     61 |   val dest_comb : term -> {Rator : term, Rand : term}
 | 
|  |     62 |   val dest_abs  : string list -> term -> {Bvar : term, Body : term} * string list
 | 
|  |     63 |   val dest_eq     : term -> {lhs : term, rhs : term}
 | 
|  |     64 |   val dest_imp    : term -> {ant : term, conseq : term}
 | 
|  |     65 |   val dest_forall : term -> {Bvar : term, Body : term}
 | 
|  |     66 |   val dest_exists : term -> {Bvar : term, Body : term}
 | 
|  |     67 |   val dest_neg    : term -> term
 | 
|  |     68 |   val dest_conj   : term -> {conj1 : term, conj2 : term}
 | 
|  |     69 |   val dest_disj   : term -> {disj1 : term, disj2 : term}
 | 
|  |     70 |   val dest_pair   : term -> {fst : term, snd : term}
 | 
|  |     71 |   val dest_pabs   : string list -> term -> {varstruct : term, body : term, used : string list}
 | 
|  |     72 | 
 | 
|  |     73 |   val lhs   : term -> term
 | 
|  |     74 |   val rhs   : term -> term
 | 
|  |     75 |   val rand  : term -> term
 | 
|  |     76 | 
 | 
|  |     77 |   (* Query routines *)
 | 
|  |     78 |   val is_imp    : term -> bool
 | 
|  |     79 |   val is_forall : term -> bool
 | 
|  |     80 |   val is_exists : term -> bool
 | 
|  |     81 |   val is_neg    : term -> bool
 | 
|  |     82 |   val is_conj   : term -> bool
 | 
|  |     83 |   val is_disj   : term -> bool
 | 
|  |     84 |   val is_pair   : term -> bool
 | 
|  |     85 |   val is_pabs   : term -> bool
 | 
|  |     86 | 
 | 
|  |     87 |   (* Construction of a term from a list of Preterms *)
 | 
|  |     88 |   val list_mk_abs    : (term list * term) -> term
 | 
|  |     89 |   val list_mk_imp    : (term list * term) -> term
 | 
|  |     90 |   val list_mk_forall : (term list * term) -> term
 | 
|  |     91 |   val list_mk_conj   : term list -> term
 | 
|  |     92 | 
 | 
|  |     93 |   (* Destructing a term to a list of Preterms *)
 | 
|  |     94 |   val strip_comb     : term -> (term * term list)
 | 
|  |     95 |   val strip_abs      : term -> (term list * term)
 | 
|  |     96 |   val strip_imp      : term -> (term list * term)
 | 
|  |     97 |   val strip_forall   : term -> (term list * term)
 | 
|  |     98 |   val strip_exists   : term -> (term list * term)
 | 
|  |     99 |   val strip_disj     : term -> term list
 | 
|  |    100 | 
 | 
|  |    101 |   (* Miscellaneous *)
 | 
|  |    102 |   val mk_vstruct : typ -> term list -> term
 | 
|  |    103 |   val gen_all    : term -> term
 | 
|  |    104 |   val find_term  : (term -> bool) -> term -> term option
 | 
|  |    105 |   val dest_relation : term -> term * term * term
 | 
|  |    106 |   val is_WFR : term -> bool
 | 
|  |    107 |   val ARB : typ -> term
 | 
|  |    108 | end;
 | 
|  |    109 | 
 | 
|  |    110 | signature DCTERM =
 | 
|  |    111 | sig
 | 
|  |    112 |   val dest_comb: cterm -> cterm * cterm
 | 
|  |    113 |   val dest_abs: string option -> cterm -> cterm * cterm
 | 
|  |    114 |   val capply: cterm -> cterm -> cterm
 | 
|  |    115 |   val cabs: cterm -> cterm -> cterm
 | 
|  |    116 |   val mk_conj: cterm * cterm -> cterm
 | 
|  |    117 |   val mk_disj: cterm * cterm -> cterm
 | 
|  |    118 |   val mk_exists: cterm * cterm -> cterm
 | 
|  |    119 |   val dest_conj: cterm -> cterm * cterm
 | 
|  |    120 |   val dest_const: cterm -> {Name: string, Ty: typ}
 | 
|  |    121 |   val dest_disj: cterm -> cterm * cterm
 | 
|  |    122 |   val dest_eq: cterm -> cterm * cterm
 | 
|  |    123 |   val dest_exists: cterm -> cterm * cterm
 | 
|  |    124 |   val dest_forall: cterm -> cterm * cterm
 | 
|  |    125 |   val dest_imp: cterm -> cterm * cterm
 | 
|  |    126 |   val dest_neg: cterm -> cterm
 | 
|  |    127 |   val dest_pair: cterm -> cterm * cterm
 | 
|  |    128 |   val dest_var: cterm -> {Name:string, Ty:typ}
 | 
|  |    129 |   val is_conj: cterm -> bool
 | 
|  |    130 |   val is_disj: cterm -> bool
 | 
|  |    131 |   val is_eq: cterm -> bool
 | 
|  |    132 |   val is_exists: cterm -> bool
 | 
|  |    133 |   val is_forall: cterm -> bool
 | 
|  |    134 |   val is_imp: cterm -> bool
 | 
|  |    135 |   val is_neg: cterm -> bool
 | 
|  |    136 |   val is_pair: cterm -> bool
 | 
|  |    137 |   val list_mk_disj: cterm list -> cterm
 | 
|  |    138 |   val strip_abs: cterm -> cterm list * cterm
 | 
|  |    139 |   val strip_comb: cterm -> cterm * cterm list
 | 
|  |    140 |   val strip_disj: cterm -> cterm list
 | 
|  |    141 |   val strip_exists: cterm -> cterm list * cterm
 | 
|  |    142 |   val strip_forall: cterm -> cterm list * cterm
 | 
|  |    143 |   val strip_imp: cterm -> cterm list * cterm
 | 
|  |    144 |   val drop_prop: cterm -> cterm
 | 
|  |    145 |   val mk_prop: cterm -> cterm
 | 
|  |    146 | end;
 | 
|  |    147 | 
 | 
|  |    148 | signature RULES =
 | 
|  |    149 | sig
 | 
|  |    150 |   val dest_thm: thm -> term list * term
 | 
|  |    151 | 
 | 
|  |    152 |   (* Inference rules *)
 | 
|  |    153 |   val REFL: cterm -> thm
 | 
|  |    154 |   val ASSUME: cterm -> thm
 | 
|  |    155 |   val MP: thm -> thm -> thm
 | 
|  |    156 |   val MATCH_MP: thm -> thm -> thm
 | 
|  |    157 |   val CONJUNCT1: thm -> thm
 | 
|  |    158 |   val CONJUNCT2: thm -> thm
 | 
|  |    159 |   val CONJUNCTS: thm -> thm list
 | 
|  |    160 |   val DISCH: cterm -> thm -> thm
 | 
|  |    161 |   val UNDISCH: thm  -> thm
 | 
|  |    162 |   val SPEC: cterm -> thm -> thm
 | 
|  |    163 |   val ISPEC: cterm -> thm -> thm
 | 
|  |    164 |   val ISPECL: cterm list -> thm -> thm
 | 
|  |    165 |   val GEN: Proof.context -> cterm -> thm -> thm
 | 
|  |    166 |   val GENL: Proof.context -> cterm list -> thm -> thm
 | 
|  |    167 |   val LIST_CONJ: thm list -> thm
 | 
|  |    168 | 
 | 
|  |    169 |   val SYM: thm -> thm
 | 
|  |    170 |   val DISCH_ALL: thm -> thm
 | 
|  |    171 |   val FILTER_DISCH_ALL: (term -> bool) -> thm -> thm
 | 
|  |    172 |   val SPEC_ALL: thm -> thm
 | 
|  |    173 |   val GEN_ALL: Proof.context -> thm -> thm
 | 
|  |    174 |   val IMP_TRANS: thm -> thm -> thm
 | 
|  |    175 |   val PROVE_HYP: thm -> thm -> thm
 | 
|  |    176 | 
 | 
|  |    177 |   val CHOOSE: Proof.context -> cterm * thm -> thm -> thm
 | 
|  |    178 |   val EXISTS: cterm * cterm -> thm -> thm
 | 
|  |    179 |   val EXISTL: cterm list -> thm -> thm
 | 
|  |    180 |   val IT_EXISTS: Proof.context -> (cterm * cterm) list -> thm -> thm
 | 
|  |    181 | 
 | 
|  |    182 |   val EVEN_ORS: thm list -> thm list
 | 
|  |    183 |   val DISJ_CASESL: thm -> thm list -> thm
 | 
|  |    184 | 
 | 
|  |    185 |   val list_beta_conv: cterm -> cterm list -> thm
 | 
|  |    186 |   val SUBS: Proof.context -> thm list -> thm -> thm
 | 
|  |    187 |   val simpl_conv: Proof.context -> thm list -> cterm -> thm
 | 
|  |    188 | 
 | 
|  |    189 |   val rbeta: thm -> thm
 | 
|  |    190 |   val tracing: bool Unsynchronized.ref
 | 
|  |    191 |   val CONTEXT_REWRITE_RULE: Proof.context ->
 | 
|  |    192 |     term * term list * thm * thm list -> thm -> thm * term list
 | 
|  |    193 |   val RIGHT_ASSOC: Proof.context -> thm -> thm
 | 
|  |    194 | 
 | 
|  |    195 |   val prove: Proof.context -> bool -> term * tactic -> thm
 | 
|  |    196 | end;
 | 
|  |    197 | 
 | 
|  |    198 | signature THRY =
 | 
|  |    199 | sig
 | 
|  |    200 |   val match_term: theory -> term -> term -> (term * term) list * (typ * typ) list
 | 
|  |    201 |   val match_type: theory -> typ -> typ -> (typ * typ) list
 | 
|  |    202 |   val typecheck: theory -> term -> cterm
 | 
|  |    203 |   (*datatype facts of various flavours*)
 | 
|  |    204 |   val match_info: theory -> string -> {constructors: term list, case_const: term} option
 | 
|  |    205 |   val induct_info: theory -> string -> {constructors: term list, nchotomy: thm} option
 | 
|  |    206 |   val extract_info: theory -> {case_congs: thm list, case_rewrites: thm list}
 | 
|  |    207 | end;
 | 
|  |    208 | 
 | 
|  |    209 | signature PRIM =
 | 
|  |    210 | sig
 | 
|  |    211 |   val trace: bool Unsynchronized.ref
 | 
|  |    212 |   val trace_thms: Proof.context -> string -> thm list -> unit
 | 
|  |    213 |   val trace_cterm: Proof.context -> string -> cterm -> unit
 | 
|  |    214 |   type pattern
 | 
|  |    215 |   val mk_functional: theory -> term list -> {functional: term, pats: pattern list}
 | 
|  |    216 |   val wfrec_definition0: string -> term -> term -> theory -> thm * theory
 | 
|  |    217 |   val post_definition: Proof.context -> thm list -> thm * pattern list ->
 | 
|  |    218 |    {rules: thm,
 | 
|  |    219 |     rows: int list,
 | 
|  |    220 |     TCs: term list list,
 | 
|  |    221 |     full_pats_TCs: (term * term list) list}
 | 
|  |    222 |   val wfrec_eqns: theory -> xstring -> thm list -> term list ->
 | 
|  |    223 |    {WFR: term,
 | 
|  |    224 |     SV: term list,
 | 
|  |    225 |     proto_def: term,
 | 
|  |    226 |     extracta: (thm * term list) list,
 | 
|  |    227 |     pats: pattern list}
 | 
|  |    228 |   val mk_induction: theory ->
 | 
|  |    229 |     {fconst: term, R: term, SV: term list, pat_TCs_list: (term * term list) list} -> thm
 | 
|  |    230 |   val postprocess: Proof.context -> bool ->
 | 
|  |    231 |     {wf_tac: tactic, terminator: tactic, simplifier: cterm -> thm} ->
 | 
|  |    232 |     {rules: thm, induction: thm, TCs: term list list} ->
 | 
|  |    233 |     {rules: thm, induction: thm, nested_tcs: thm list}
 | 
|  |    234 | end;
 | 
|  |    235 | 
 | 
|  |    236 | signature TFL =
 | 
|  |    237 | sig
 | 
|  |    238 |   val define_i: bool -> thm list -> thm list -> xstring -> term -> term list -> Proof.context ->
 | 
|  |    239 |     {lhs: term, rules: (thm * int) list, induct: thm, tcs: term list} * Proof.context
 | 
|  |    240 |   val define: bool -> thm list -> thm list -> xstring -> string -> string list -> Proof.context ->
 | 
|  |    241 |     {lhs: term, rules: (thm * int) list, induct: thm, tcs: term list} * Proof.context
 | 
|  |    242 | end;
 | 
|  |    243 | 
 | 
|  |    244 | signature OLD_RECDEF =
 | 
|  |    245 | sig
 | 
|  |    246 |   val get_recdef: theory -> string
 | 
|  |    247 |     -> {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list} option
 | 
|  |    248 |   val get_hints: Proof.context -> {simps: thm list, congs: (string * thm) list, wfs: thm list}
 | 
|  |    249 |   val simp_add: attribute
 | 
|  |    250 |   val simp_del: attribute
 | 
|  |    251 |   val cong_add: attribute
 | 
|  |    252 |   val cong_del: attribute
 | 
|  |    253 |   val wf_add: attribute
 | 
|  |    254 |   val wf_del: attribute
 | 
|  |    255 |   val add_recdef: bool -> xstring -> string -> ((binding * string) * Token.src list) list ->
 | 
|  |    256 |     Token.src option -> theory -> theory
 | 
|  |    257 |       * {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list}
 | 
|  |    258 |   val add_recdef_i: bool -> xstring -> term -> ((binding * term) * attribute list) list ->
 | 
|  |    259 |     theory -> theory * {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list}
 | 
|  |    260 | end;
 | 
|  |    261 | 
 | 
|  |    262 | structure Old_Recdef: OLD_RECDEF =
 | 
|  |    263 | struct
 | 
|  |    264 | 
 | 
|  |    265 | (*** extra case splitting for TFL ***)
 | 
|  |    266 | 
 | 
|  |    267 | structure CaseSplit: CASE_SPLIT =
 | 
|  |    268 | struct
 | 
|  |    269 | 
 | 
|  |    270 | (* make a casethm from an induction thm *)
 | 
|  |    271 | val cases_thm_of_induct_thm =
 | 
|  |    272 |      Seq.hd o (ALLGOALS (fn i => REPEAT (etac Drule.thin_rl i)));
 | 
|  |    273 | 
 | 
|  |    274 | (* get the case_thm (my version) from a type *)
 | 
|  |    275 | fun case_thm_of_ty thy ty  =
 | 
|  |    276 |     let
 | 
|  |    277 |       val ty_str = case ty of
 | 
|  |    278 |                      Type(ty_str, _) => ty_str
 | 
|  |    279 |                    | TFree(s,_)  => error ("Free type: " ^ s)
 | 
| 60521 |    280 |                    | TVar((s,_),_) => error ("Free variable: " ^ s)
 | 
| 60520 |    281 |       val {induct, ...} = BNF_LFP_Compat.the_info thy [BNF_LFP_Compat.Keep_Nesting] ty_str
 | 
|  |    282 |     in
 | 
|  |    283 |       cases_thm_of_induct_thm induct
 | 
|  |    284 |     end;
 | 
|  |    285 | 
 | 
|  |    286 | 
 | 
|  |    287 | (* for use when there are no prems to the subgoal *)
 | 
|  |    288 | (* does a case split on the given variable *)
 | 
|  |    289 | fun mk_casesplit_goal_thm ctxt (vstr,ty) gt =
 | 
|  |    290 |     let
 | 
|  |    291 |       val thy = Proof_Context.theory_of ctxt;
 | 
|  |    292 | 
 | 
|  |    293 |       val x = Free(vstr,ty);
 | 
|  |    294 |       val abst = Abs(vstr, ty, Term.abstract_over (x, gt));
 | 
|  |    295 | 
 | 
|  |    296 |       val case_thm = case_thm_of_ty thy ty;
 | 
|  |    297 | 
 | 
|  |    298 |       val abs_ct = Thm.cterm_of ctxt abst;
 | 
|  |    299 |       val free_ct = Thm.cterm_of ctxt x;
 | 
|  |    300 | 
 | 
|  |    301 |       val (Pv, Dv, type_insts) =
 | 
|  |    302 |           case (Thm.concl_of case_thm) of
 | 
| 60521 |    303 |             (_ $ (Pv $ (Dv as Var(_, Dty)))) =>
 | 
| 60520 |    304 |             (Pv, Dv,
 | 
|  |    305 |              Sign.typ_match thy (Dty, ty) Vartab.empty)
 | 
|  |    306 |           | _ => error "not a valid case thm";
 | 
|  |    307 |       val type_cinsts = map (fn (ixn, (S, T)) => apply2 (Thm.ctyp_of ctxt) (TVar (ixn, S), T))
 | 
|  |    308 |         (Vartab.dest type_insts);
 | 
|  |    309 |       val cPv = Thm.cterm_of ctxt (Envir.subst_term_types type_insts Pv);
 | 
|  |    310 |       val cDv = Thm.cterm_of ctxt (Envir.subst_term_types type_insts Dv);
 | 
|  |    311 |     in
 | 
|  |    312 |       Conv.fconv_rule Drule.beta_eta_conversion
 | 
|  |    313 |          (case_thm
 | 
|  |    314 |             |> Thm.instantiate (type_cinsts, [])
 | 
|  |    315 |             |> Thm.instantiate ([], [(cPv, abs_ct), (cDv, free_ct)]))
 | 
|  |    316 |     end;
 | 
|  |    317 | 
 | 
|  |    318 | 
 | 
|  |    319 | (* the find_XXX_split functions are simply doing a lightwieght (I
 | 
|  |    320 | think) term matching equivalent to find where to do the next split *)
 | 
|  |    321 | 
 | 
|  |    322 | (* assuming two twems are identical except for a free in one at a
 | 
|  |    323 | subterm, or constant in another, ie assume that one term is a plit of
 | 
|  |    324 | another, then gives back the free variable that has been split. *)
 | 
|  |    325 | exception find_split_exp of string
 | 
|  |    326 | fun find_term_split (Free v, _ $ _) = SOME v
 | 
|  |    327 |   | find_term_split (Free v, Const _) = SOME v
 | 
|  |    328 |   | find_term_split (Free v, Abs _) = SOME v (* do we really want this case? *)
 | 
| 60521 |    329 |   | find_term_split (Free _, Var _) = NONE (* keep searching *)
 | 
| 60520 |    330 |   | find_term_split (a $ b, a2 $ b2) =
 | 
|  |    331 |     (case find_term_split (a, a2) of
 | 
|  |    332 |        NONE => find_term_split (b,b2)
 | 
|  |    333 |      | vopt => vopt)
 | 
| 60521 |    334 |   | find_term_split (Abs(_,_,t1), Abs(_,_,t2)) =
 | 
| 60520 |    335 |     find_term_split (t1, t2)
 | 
| 60521 |    336 |   | find_term_split (Const (x,_), Const(x2,_)) =
 | 
| 60520 |    337 |     if x = x2 then NONE else (* keep searching *)
 | 
|  |    338 |     raise find_split_exp (* stop now *)
 | 
|  |    339 |             "Terms are not identical upto a free varaible! (Consts)"
 | 
|  |    340 |   | find_term_split (Bound i, Bound j) =
 | 
|  |    341 |     if i = j then NONE else (* keep searching *)
 | 
|  |    342 |     raise find_split_exp (* stop now *)
 | 
|  |    343 |             "Terms are not identical upto a free varaible! (Bound)"
 | 
|  |    344 |   | find_term_split _ =
 | 
|  |    345 |     raise find_split_exp (* stop now *)
 | 
|  |    346 |             "Terms are not identical upto a free varaible! (Other)";
 | 
|  |    347 | 
 | 
|  |    348 | (* assume that "splitth" is a case split form of subgoal i of "genth",
 | 
|  |    349 | then look for a free variable to split, breaking the subgoal closer to
 | 
|  |    350 | splitth. *)
 | 
|  |    351 | fun find_thm_split splitth i genth =
 | 
|  |    352 |     find_term_split (Logic.get_goal (Thm.prop_of genth) i,
 | 
|  |    353 |                      Thm.concl_of splitth) handle find_split_exp _ => NONE;
 | 
|  |    354 | 
 | 
|  |    355 | (* as above but searches "splitths" for a theorem that suggest a case split *)
 | 
|  |    356 | fun find_thms_split splitths i genth =
 | 
|  |    357 |     Library.get_first (fn sth => find_thm_split sth i genth) splitths;
 | 
|  |    358 | 
 | 
|  |    359 | 
 | 
|  |    360 | (* split the subgoal i of "genth" until we get to a member of
 | 
|  |    361 | splitths. Assumes that genth will be a general form of splitths, that
 | 
|  |    362 | can be case-split, as needed. Otherwise fails. Note: We assume that
 | 
|  |    363 | all of "splitths" are split to the same level, and thus it doesn't
 | 
|  |    364 | matter which one we choose to look for the next split. Simply add
 | 
|  |    365 | search on splitthms and split variable, to change this.  *)
 | 
|  |    366 | (* Note: possible efficiency measure: when a case theorem is no longer
 | 
|  |    367 | useful, drop it? *)
 | 
|  |    368 | (* Note: This should not be a separate tactic but integrated into the
 | 
|  |    369 | case split done during recdef's case analysis, this would avoid us
 | 
|  |    370 | having to (re)search for variables to split. *)
 | 
|  |    371 | fun splitto ctxt splitths genth =
 | 
|  |    372 |     let
 | 
|  |    373 |       val _ = not (null splitths) orelse error "splitto: no given splitths";
 | 
|  |    374 | 
 | 
|  |    375 |       (* check if we are a member of splitths - FIXME: quicker and
 | 
|  |    376 |       more flexible with discrim net. *)
 | 
|  |    377 |       fun solve_by_splitth th split =
 | 
|  |    378 |         Thm.biresolution (SOME ctxt) false [(false,split)] 1 th;
 | 
|  |    379 | 
 | 
|  |    380 |       fun split th =
 | 
|  |    381 |         (case find_thms_split splitths 1 th of
 | 
|  |    382 |           NONE =>
 | 
|  |    383 |            (writeln (cat_lines
 | 
|  |    384 |             (["th:", Display.string_of_thm ctxt th, "split ths:"] @
 | 
|  |    385 |               map (Display.string_of_thm ctxt) splitths @ ["\n--"]));
 | 
|  |    386 |             error "splitto: cannot find variable to split on")
 | 
|  |    387 |         | SOME v =>
 | 
|  |    388 |             let
 | 
|  |    389 |               val gt = HOLogic.dest_Trueprop (#1 (Logic.dest_implies (Thm.prop_of th)));
 | 
|  |    390 |               val split_thm = mk_casesplit_goal_thm ctxt v gt;
 | 
|  |    391 |               val (subthms, expf) = IsaND.fixed_subgoal_thms ctxt split_thm;
 | 
|  |    392 |             in
 | 
|  |    393 |               expf (map recsplitf subthms)
 | 
|  |    394 |             end)
 | 
|  |    395 | 
 | 
|  |    396 |       and recsplitf th =
 | 
|  |    397 |         (* note: multiple unifiers! we only take the first element,
 | 
|  |    398 |            probably fine -- there is probably only one anyway. *)
 | 
|  |    399 |         (case get_first (Seq.pull o solve_by_splitth th) splitths of
 | 
|  |    400 |           NONE => split th
 | 
|  |    401 |         | SOME (solved_th, _) => solved_th);
 | 
|  |    402 |     in
 | 
|  |    403 |       recsplitf genth
 | 
|  |    404 |     end;
 | 
|  |    405 | 
 | 
|  |    406 | end;
 | 
|  |    407 | 
 | 
|  |    408 | 
 | 
| 60521 |    409 | 
 | 
| 60520 |    410 | (*** basic utilities ***)
 | 
|  |    411 | 
 | 
|  |    412 | structure Utils: UTILS =
 | 
|  |    413 | struct
 | 
|  |    414 | 
 | 
|  |    415 | (*standard exception for TFL*)
 | 
|  |    416 | exception ERR of {module: string, func: string, mesg: string};
 | 
|  |    417 | 
 | 
|  |    418 | fun UTILS_ERR func mesg = ERR {module = "Utils", func = func, mesg = mesg};
 | 
|  |    419 | 
 | 
|  |    420 | 
 | 
| 60521 |    421 | fun end_itlist _ [] = raise (UTILS_ERR "end_itlist" "list too short")
 | 
|  |    422 |   | end_itlist _ [x] = x
 | 
| 60520 |    423 |   | end_itlist f (x :: xs) = f x (end_itlist f xs);
 | 
|  |    424 | 
 | 
|  |    425 | fun itlist2 f L1 L2 base_value =
 | 
|  |    426 |  let fun it ([],[]) = base_value
 | 
|  |    427 |        | it ((a::rst1),(b::rst2)) = f a b (it (rst1,rst2))
 | 
|  |    428 |        | it _ = raise UTILS_ERR "itlist2" "different length lists"
 | 
|  |    429 |  in  it (L1,L2)
 | 
|  |    430 |  end;
 | 
|  |    431 | 
 | 
|  |    432 | fun pluck p  =
 | 
|  |    433 |   let fun remv ([],_) = raise UTILS_ERR "pluck" "item not found"
 | 
|  |    434 |         | remv (h::t, A) = if p h then (h, rev A @ t) else remv (t,h::A)
 | 
|  |    435 |   in fn L => remv(L,[])
 | 
|  |    436 |   end;
 | 
|  |    437 | 
 | 
|  |    438 | fun take f =
 | 
| 60521 |    439 |   let fun grab(0, _) = []
 | 
| 60520 |    440 |         | grab(n, x::rst) = f x::grab(n-1,rst)
 | 
|  |    441 |   in grab
 | 
|  |    442 |   end;
 | 
|  |    443 | 
 | 
|  |    444 | fun zip3 [][][] = []
 | 
|  |    445 |   | zip3 (x::l1) (y::l2) (z::l3) = (x,y,z)::zip3 l1 l2 l3
 | 
|  |    446 |   | zip3 _ _ _ = raise UTILS_ERR "zip3" "different lengths";
 | 
|  |    447 | 
 | 
|  |    448 | 
 | 
|  |    449 | end;
 | 
|  |    450 | 
 | 
|  |    451 | 
 | 
| 60521 |    452 | 
 | 
| 60520 |    453 | (*** emulation of HOL's abstract syntax functions ***)
 | 
|  |    454 | 
 | 
|  |    455 | structure USyntax: USYNTAX =
 | 
|  |    456 | struct
 | 
|  |    457 | 
 | 
|  |    458 | infix 4 ##;
 | 
|  |    459 | 
 | 
|  |    460 | fun USYN_ERR func mesg = Utils.ERR {module = "USyntax", func = func, mesg = mesg};
 | 
|  |    461 | 
 | 
|  |    462 | 
 | 
|  |    463 | (*---------------------------------------------------------------------------
 | 
|  |    464 |  *
 | 
|  |    465 |  *                            Types
 | 
|  |    466 |  *
 | 
|  |    467 |  *---------------------------------------------------------------------------*)
 | 
|  |    468 | val mk_prim_vartype = TVar;
 | 
|  |    469 | fun mk_vartype s = mk_prim_vartype ((s, 0), @{sort type});
 | 
|  |    470 | 
 | 
|  |    471 | (* But internally, it's useful *)
 | 
|  |    472 | fun dest_vtype (TVar x) = x
 | 
|  |    473 |   | dest_vtype _ = raise USYN_ERR "dest_vtype" "not a flexible type variable";
 | 
|  |    474 | 
 | 
|  |    475 | val is_vartype = can dest_vtype;
 | 
|  |    476 | 
 | 
|  |    477 | val type_vars  = map mk_prim_vartype o Misc_Legacy.typ_tvars
 | 
|  |    478 | fun type_varsl L = distinct (op =) (fold (curry op @ o type_vars) L []);
 | 
|  |    479 | 
 | 
|  |    480 | val alpha  = mk_vartype "'a"
 | 
|  |    481 | 
 | 
|  |    482 | val strip_prod_type = HOLogic.flatten_tupleT;
 | 
|  |    483 | 
 | 
|  |    484 | 
 | 
|  |    485 | 
 | 
|  |    486 | (*---------------------------------------------------------------------------
 | 
|  |    487 |  *
 | 
|  |    488 |  *                              Terms
 | 
|  |    489 |  *
 | 
|  |    490 |  *---------------------------------------------------------------------------*)
 | 
|  |    491 | 
 | 
|  |    492 | (* Free variables, in order of occurrence, from left to right in the
 | 
|  |    493 |  * syntax tree. *)
 | 
|  |    494 | fun free_vars_lr tm =
 | 
|  |    495 |   let fun memb x = let fun m[] = false | m(y::rst) = (x=y)orelse m rst in m end
 | 
|  |    496 |       fun add (t, frees) = case t of
 | 
|  |    497 |             Free   _ => if (memb t frees) then frees else t::frees
 | 
|  |    498 |           | Abs (_,_,body) => add(body,frees)
 | 
|  |    499 |           | f$t =>  add(t, add(f, frees))
 | 
|  |    500 |           | _ => frees
 | 
|  |    501 |   in rev(add(tm,[]))
 | 
|  |    502 |   end;
 | 
|  |    503 | 
 | 
|  |    504 | 
 | 
|  |    505 | 
 | 
|  |    506 | val type_vars_in_term = map mk_prim_vartype o Misc_Legacy.term_tvars;
 | 
|  |    507 | 
 | 
|  |    508 | 
 | 
|  |    509 | 
 | 
|  |    510 | (* Prelogic *)
 | 
|  |    511 | fun dest_tybinding (v,ty) = (#1(dest_vtype v),ty)
 | 
|  |    512 | fun inst theta = subst_vars (map dest_tybinding theta,[])
 | 
|  |    513 | 
 | 
|  |    514 | 
 | 
|  |    515 | (* Construction routines *)
 | 
|  |    516 | 
 | 
|  |    517 | fun mk_abs{Bvar as Var((s,_),ty),Body}  = Abs(s,ty,abstract_over(Bvar,Body))
 | 
|  |    518 |   | mk_abs{Bvar as Free(s,ty),Body}  = Abs(s,ty,abstract_over(Bvar,Body))
 | 
|  |    519 |   | mk_abs _ = raise USYN_ERR "mk_abs" "Bvar is not a variable";
 | 
|  |    520 | 
 | 
|  |    521 | 
 | 
|  |    522 | fun mk_imp{ant,conseq} =
 | 
|  |    523 |    let val c = Const(@{const_name HOL.implies},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
 | 
|  |    524 |    in list_comb(c,[ant,conseq])
 | 
|  |    525 |    end;
 | 
|  |    526 | 
 | 
|  |    527 | fun mk_select (r as {Bvar,Body}) =
 | 
|  |    528 |   let val ty = type_of Bvar
 | 
|  |    529 |       val c = Const(@{const_name Eps},(ty --> HOLogic.boolT) --> ty)
 | 
|  |    530 |   in list_comb(c,[mk_abs r])
 | 
|  |    531 |   end;
 | 
|  |    532 | 
 | 
|  |    533 | fun mk_forall (r as {Bvar,Body}) =
 | 
|  |    534 |   let val ty = type_of Bvar
 | 
|  |    535 |       val c = Const(@{const_name All},(ty --> HOLogic.boolT) --> HOLogic.boolT)
 | 
|  |    536 |   in list_comb(c,[mk_abs r])
 | 
|  |    537 |   end;
 | 
|  |    538 | 
 | 
|  |    539 | fun mk_exists (r as {Bvar,Body}) =
 | 
|  |    540 |   let val ty = type_of Bvar
 | 
|  |    541 |       val c = Const(@{const_name Ex},(ty --> HOLogic.boolT) --> HOLogic.boolT)
 | 
|  |    542 |   in list_comb(c,[mk_abs r])
 | 
|  |    543 |   end;
 | 
|  |    544 | 
 | 
|  |    545 | 
 | 
|  |    546 | fun mk_conj{conj1,conj2} =
 | 
|  |    547 |    let val c = Const(@{const_name HOL.conj},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
 | 
|  |    548 |    in list_comb(c,[conj1,conj2])
 | 
|  |    549 |    end;
 | 
|  |    550 | 
 | 
|  |    551 | fun mk_disj{disj1,disj2} =
 | 
|  |    552 |    let val c = Const(@{const_name HOL.disj},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
 | 
|  |    553 |    in list_comb(c,[disj1,disj2])
 | 
|  |    554 |    end;
 | 
|  |    555 | 
 | 
|  |    556 | fun prod_ty ty1 ty2 = HOLogic.mk_prodT (ty1,ty2);
 | 
|  |    557 | 
 | 
|  |    558 | local
 | 
|  |    559 | fun mk_uncurry (xt, yt, zt) =
 | 
|  |    560 |     Const(@{const_name case_prod}, (xt --> yt --> zt) --> prod_ty xt yt --> zt)
 | 
|  |    561 | fun dest_pair(Const(@{const_name Pair},_) $ M $ N) = {fst=M, snd=N}
 | 
|  |    562 |   | dest_pair _ = raise USYN_ERR "dest_pair" "not a pair"
 | 
|  |    563 | fun is_var (Var _) = true | is_var (Free _) = true | is_var _ = false
 | 
|  |    564 | in
 | 
|  |    565 | fun mk_pabs{varstruct,body} =
 | 
|  |    566 |  let fun mpa (varstruct, body) =
 | 
|  |    567 |        if is_var varstruct
 | 
|  |    568 |        then mk_abs {Bvar = varstruct, Body = body}
 | 
|  |    569 |        else let val {fst, snd} = dest_pair varstruct
 | 
|  |    570 |             in mk_uncurry (type_of fst, type_of snd, type_of body) $
 | 
|  |    571 |                mpa (fst, mpa (snd, body))
 | 
|  |    572 |             end
 | 
|  |    573 |  in mpa (varstruct, body) end
 | 
|  |    574 |  handle TYPE _ => raise USYN_ERR "mk_pabs" "";
 | 
|  |    575 | end;
 | 
|  |    576 | 
 | 
|  |    577 | (* Destruction routines *)
 | 
|  |    578 | 
 | 
|  |    579 | datatype lambda = VAR   of {Name : string, Ty : typ}
 | 
|  |    580 |                 | CONST of {Name : string, Ty : typ}
 | 
|  |    581 |                 | COMB  of {Rator: term, Rand : term}
 | 
|  |    582 |                 | LAMB  of {Bvar : term, Body : term};
 | 
|  |    583 | 
 | 
|  |    584 | 
 | 
| 60521 |    585 | fun dest_term(Var((s,_),ty)) = VAR{Name = s, Ty = ty}
 | 
| 60520 |    586 |   | dest_term(Free(s,ty))    = VAR{Name = s, Ty = ty}
 | 
|  |    587 |   | dest_term(Const(s,ty))   = CONST{Name = s, Ty = ty}
 | 
|  |    588 |   | dest_term(M$N)           = COMB{Rator=M,Rand=N}
 | 
|  |    589 |   | dest_term(Abs(s,ty,M))   = let  val v = Free(s,ty)
 | 
|  |    590 |                                in LAMB{Bvar = v, Body = Term.betapply (M,v)}
 | 
|  |    591 |                                end
 | 
|  |    592 |   | dest_term(Bound _)       = raise USYN_ERR "dest_term" "Bound";
 | 
|  |    593 | 
 | 
|  |    594 | fun dest_const(Const(s,ty)) = {Name = s, Ty = ty}
 | 
|  |    595 |   | dest_const _ = raise USYN_ERR "dest_const" "not a constant";
 | 
|  |    596 | 
 | 
|  |    597 | fun dest_comb(t1 $ t2) = {Rator = t1, Rand = t2}
 | 
|  |    598 |   | dest_comb _ =  raise USYN_ERR "dest_comb" "not a comb";
 | 
|  |    599 | 
 | 
| 60524 |    600 | fun dest_abs used (a as Abs(s, ty, _)) =
 | 
| 60520 |    601 |      let
 | 
|  |    602 |        val s' = singleton (Name.variant_list used) s;
 | 
|  |    603 |        val v = Free(s', ty);
 | 
|  |    604 |      in ({Bvar = v, Body = Term.betapply (a,v)}, s'::used)
 | 
|  |    605 |      end
 | 
|  |    606 |   | dest_abs _ _ =  raise USYN_ERR "dest_abs" "not an abstraction";
 | 
|  |    607 | 
 | 
|  |    608 | fun dest_eq(Const(@{const_name HOL.eq},_) $ M $ N) = {lhs=M, rhs=N}
 | 
|  |    609 |   | dest_eq _ = raise USYN_ERR "dest_eq" "not an equality";
 | 
|  |    610 | 
 | 
|  |    611 | fun dest_imp(Const(@{const_name HOL.implies},_) $ M $ N) = {ant=M, conseq=N}
 | 
|  |    612 |   | dest_imp _ = raise USYN_ERR "dest_imp" "not an implication";
 | 
|  |    613 | 
 | 
|  |    614 | fun dest_forall(Const(@{const_name All},_) $ (a as Abs _)) = fst (dest_abs [] a)
 | 
|  |    615 |   | dest_forall _ = raise USYN_ERR "dest_forall" "not a forall";
 | 
|  |    616 | 
 | 
|  |    617 | fun dest_exists(Const(@{const_name Ex},_) $ (a as Abs _)) = fst (dest_abs [] a)
 | 
|  |    618 |   | dest_exists _ = raise USYN_ERR "dest_exists" "not an existential";
 | 
|  |    619 | 
 | 
|  |    620 | fun dest_neg(Const(@{const_name Not},_) $ M) = M
 | 
|  |    621 |   | dest_neg _ = raise USYN_ERR "dest_neg" "not a negation";
 | 
|  |    622 | 
 | 
|  |    623 | fun dest_conj(Const(@{const_name HOL.conj},_) $ M $ N) = {conj1=M, conj2=N}
 | 
|  |    624 |   | dest_conj _ = raise USYN_ERR "dest_conj" "not a conjunction";
 | 
|  |    625 | 
 | 
|  |    626 | fun dest_disj(Const(@{const_name HOL.disj},_) $ M $ N) = {disj1=M, disj2=N}
 | 
|  |    627 |   | dest_disj _ = raise USYN_ERR "dest_disj" "not a disjunction";
 | 
|  |    628 | 
 | 
|  |    629 | fun mk_pair{fst,snd} =
 | 
|  |    630 |    let val ty1 = type_of fst
 | 
|  |    631 |        val ty2 = type_of snd
 | 
|  |    632 |        val c = Const(@{const_name Pair},ty1 --> ty2 --> prod_ty ty1 ty2)
 | 
|  |    633 |    in list_comb(c,[fst,snd])
 | 
|  |    634 |    end;
 | 
|  |    635 | 
 | 
|  |    636 | fun dest_pair(Const(@{const_name Pair},_) $ M $ N) = {fst=M, snd=N}
 | 
|  |    637 |   | dest_pair _ = raise USYN_ERR "dest_pair" "not a pair";
 | 
|  |    638 | 
 | 
|  |    639 | 
 | 
|  |    640 | local  fun ucheck t = (if #Name (dest_const t) = @{const_name case_prod} then t
 | 
|  |    641 |                        else raise Match)
 | 
|  |    642 | in
 | 
|  |    643 | fun dest_pabs used tm =
 | 
|  |    644 |    let val ({Bvar,Body}, used') = dest_abs used tm
 | 
|  |    645 |    in {varstruct = Bvar, body = Body, used = used'}
 | 
|  |    646 |    end handle Utils.ERR _ =>
 | 
|  |    647 |           let val {Rator,Rand} = dest_comb tm
 | 
|  |    648 |               val _ = ucheck Rator
 | 
|  |    649 |               val {varstruct = lv, body, used = used'} = dest_pabs used Rand
 | 
|  |    650 |               val {varstruct = rv, body, used = used''} = dest_pabs used' body
 | 
|  |    651 |           in {varstruct = mk_pair {fst = lv, snd = rv}, body = body, used = used''}
 | 
|  |    652 |           end
 | 
|  |    653 | end;
 | 
|  |    654 | 
 | 
|  |    655 | 
 | 
|  |    656 | val lhs   = #lhs o dest_eq
 | 
|  |    657 | val rhs   = #rhs o dest_eq
 | 
|  |    658 | val rand  = #Rand o dest_comb
 | 
|  |    659 | 
 | 
|  |    660 | 
 | 
|  |    661 | (* Query routines *)
 | 
|  |    662 | val is_imp    = can dest_imp
 | 
|  |    663 | val is_forall = can dest_forall
 | 
|  |    664 | val is_exists = can dest_exists
 | 
|  |    665 | val is_neg    = can dest_neg
 | 
|  |    666 | val is_conj   = can dest_conj
 | 
|  |    667 | val is_disj   = can dest_disj
 | 
|  |    668 | val is_pair   = can dest_pair
 | 
|  |    669 | val is_pabs   = can (dest_pabs [])
 | 
|  |    670 | 
 | 
|  |    671 | 
 | 
|  |    672 | (* Construction of a cterm from a list of Terms *)
 | 
|  |    673 | 
 | 
|  |    674 | fun list_mk_abs(L,tm) = fold_rev (fn v => fn M => mk_abs{Bvar=v, Body=M}) L tm;
 | 
|  |    675 | 
 | 
|  |    676 | (* These others are almost never used *)
 | 
|  |    677 | fun list_mk_imp(A,c) = fold_rev (fn a => fn tm => mk_imp{ant=a,conseq=tm}) A c;
 | 
|  |    678 | fun list_mk_forall(V,t) = fold_rev (fn v => fn b => mk_forall{Bvar=v, Body=b})V t;
 | 
|  |    679 | val list_mk_conj = Utils.end_itlist(fn c1 => fn tm => mk_conj{conj1=c1, conj2=tm})
 | 
|  |    680 | 
 | 
|  |    681 | 
 | 
|  |    682 | (* Need to reverse? *)
 | 
|  |    683 | fun gen_all tm = list_mk_forall(Misc_Legacy.term_frees tm, tm);
 | 
|  |    684 | 
 | 
|  |    685 | (* Destructing a cterm to a list of Terms *)
 | 
|  |    686 | fun strip_comb tm =
 | 
|  |    687 |    let fun dest(M$N, A) = dest(M, N::A)
 | 
|  |    688 |          | dest x = x
 | 
|  |    689 |    in dest(tm,[])
 | 
|  |    690 |    end;
 | 
|  |    691 | 
 | 
|  |    692 | fun strip_abs(tm as Abs _) =
 | 
|  |    693 |        let val ({Bvar,Body}, _) = dest_abs [] tm
 | 
|  |    694 |            val (bvs, core) = strip_abs Body
 | 
|  |    695 |        in (Bvar::bvs, core)
 | 
|  |    696 |        end
 | 
|  |    697 |   | strip_abs M = ([],M);
 | 
|  |    698 | 
 | 
|  |    699 | 
 | 
|  |    700 | fun strip_imp fm =
 | 
|  |    701 |    if (is_imp fm)
 | 
|  |    702 |    then let val {ant,conseq} = dest_imp fm
 | 
|  |    703 |             val (was,wb) = strip_imp conseq
 | 
|  |    704 |         in ((ant::was), wb)
 | 
|  |    705 |         end
 | 
|  |    706 |    else ([],fm);
 | 
|  |    707 | 
 | 
|  |    708 | fun strip_forall fm =
 | 
|  |    709 |    if (is_forall fm)
 | 
|  |    710 |    then let val {Bvar,Body} = dest_forall fm
 | 
|  |    711 |             val (bvs,core) = strip_forall Body
 | 
|  |    712 |         in ((Bvar::bvs), core)
 | 
|  |    713 |         end
 | 
|  |    714 |    else ([],fm);
 | 
|  |    715 | 
 | 
|  |    716 | 
 | 
|  |    717 | fun strip_exists fm =
 | 
|  |    718 |    if (is_exists fm)
 | 
|  |    719 |    then let val {Bvar, Body} = dest_exists fm
 | 
|  |    720 |             val (bvs,core) = strip_exists Body
 | 
|  |    721 |         in (Bvar::bvs, core)
 | 
|  |    722 |         end
 | 
|  |    723 |    else ([],fm);
 | 
|  |    724 | 
 | 
|  |    725 | fun strip_disj w =
 | 
|  |    726 |    if (is_disj w)
 | 
|  |    727 |    then let val {disj1,disj2} = dest_disj w
 | 
|  |    728 |         in (strip_disj disj1@strip_disj disj2)
 | 
|  |    729 |         end
 | 
|  |    730 |    else [w];
 | 
|  |    731 | 
 | 
|  |    732 | 
 | 
|  |    733 | (* Miscellaneous *)
 | 
|  |    734 | 
 | 
|  |    735 | fun mk_vstruct ty V =
 | 
|  |    736 |   let fun follow_prod_type (Type(@{type_name Product_Type.prod},[ty1,ty2])) vs =
 | 
|  |    737 |               let val (ltm,vs1) = follow_prod_type ty1 vs
 | 
|  |    738 |                   val (rtm,vs2) = follow_prod_type ty2 vs1
 | 
|  |    739 |               in (mk_pair{fst=ltm, snd=rtm}, vs2) end
 | 
|  |    740 |         | follow_prod_type _ (v::vs) = (v,vs)
 | 
|  |    741 |   in #1 (follow_prod_type ty V)  end;
 | 
|  |    742 | 
 | 
|  |    743 | 
 | 
|  |    744 | (* Search a term for a sub-term satisfying the predicate p. *)
 | 
|  |    745 | fun find_term p =
 | 
|  |    746 |    let fun find tm =
 | 
|  |    747 |       if (p tm) then SOME tm
 | 
|  |    748 |       else case tm of
 | 
|  |    749 |           Abs(_,_,body) => find body
 | 
|  |    750 |         | (t$u)         => (case find t of NONE => find u | some => some)
 | 
|  |    751 |         | _             => NONE
 | 
|  |    752 |    in find
 | 
|  |    753 |    end;
 | 
|  |    754 | 
 | 
|  |    755 | fun dest_relation tm =
 | 
|  |    756 |    if (type_of tm = HOLogic.boolT)
 | 
|  |    757 |    then let val (Const(@{const_name Set.member},_) $ (Const(@{const_name Pair},_)$y$x) $ R) = tm
 | 
|  |    758 |         in (R,y,x)
 | 
|  |    759 |         end handle Bind => raise USYN_ERR "dest_relation" "unexpected term structure"
 | 
|  |    760 |    else raise USYN_ERR "dest_relation" "not a boolean term";
 | 
|  |    761 | 
 | 
|  |    762 | fun is_WFR (Const(@{const_name Wellfounded.wf},_)$_) = true
 | 
|  |    763 |   | is_WFR _                 = false;
 | 
|  |    764 | 
 | 
|  |    765 | fun ARB ty = mk_select{Bvar=Free("v",ty),
 | 
|  |    766 |                        Body=Const(@{const_name True},HOLogic.boolT)};
 | 
|  |    767 | 
 | 
|  |    768 | end;
 | 
|  |    769 | 
 | 
|  |    770 | 
 | 
| 60521 |    771 | 
 | 
| 60520 |    772 | (*** derived cterm destructors ***)
 | 
|  |    773 | 
 | 
|  |    774 | structure Dcterm: DCTERM =
 | 
|  |    775 | struct
 | 
|  |    776 | 
 | 
|  |    777 | fun ERR func mesg = Utils.ERR {module = "Dcterm", func = func, mesg = mesg};
 | 
|  |    778 | 
 | 
|  |    779 | 
 | 
|  |    780 | fun dest_comb t = Thm.dest_comb t
 | 
|  |    781 |   handle CTERM (msg, _) => raise ERR "dest_comb" msg;
 | 
|  |    782 | 
 | 
|  |    783 | fun dest_abs a t = Thm.dest_abs a t
 | 
|  |    784 |   handle CTERM (msg, _) => raise ERR "dest_abs" msg;
 | 
|  |    785 | 
 | 
|  |    786 | fun capply t u = Thm.apply t u
 | 
|  |    787 |   handle CTERM (msg, _) => raise ERR "capply" msg;
 | 
|  |    788 | 
 | 
|  |    789 | fun cabs a t = Thm.lambda a t
 | 
|  |    790 |   handle CTERM (msg, _) => raise ERR "cabs" msg;
 | 
|  |    791 | 
 | 
|  |    792 | 
 | 
|  |    793 | (*---------------------------------------------------------------------------
 | 
|  |    794 |  * Some simple constructor functions.
 | 
|  |    795 |  *---------------------------------------------------------------------------*)
 | 
|  |    796 | 
 | 
|  |    797 | val mk_hol_const = Thm.cterm_of @{theory_context HOL} o Const;
 | 
|  |    798 | 
 | 
|  |    799 | fun mk_exists (r as (Bvar, Body)) =
 | 
|  |    800 |   let val ty = Thm.typ_of_cterm Bvar
 | 
|  |    801 |       val c = mk_hol_const(@{const_name Ex}, (ty --> HOLogic.boolT) --> HOLogic.boolT)
 | 
|  |    802 |   in capply c (uncurry cabs r) end;
 | 
|  |    803 | 
 | 
|  |    804 | 
 | 
|  |    805 | local val c = mk_hol_const(@{const_name HOL.conj}, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
 | 
|  |    806 | in fun mk_conj(conj1,conj2) = capply (capply c conj1) conj2
 | 
|  |    807 | end;
 | 
|  |    808 | 
 | 
|  |    809 | local val c = mk_hol_const(@{const_name HOL.disj}, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
 | 
|  |    810 | in fun mk_disj(disj1,disj2) = capply (capply c disj1) disj2
 | 
|  |    811 | end;
 | 
|  |    812 | 
 | 
|  |    813 | 
 | 
|  |    814 | (*---------------------------------------------------------------------------
 | 
|  |    815 |  * The primitives.
 | 
|  |    816 |  *---------------------------------------------------------------------------*)
 | 
|  |    817 | fun dest_const ctm =
 | 
|  |    818 |    (case Thm.term_of ctm
 | 
|  |    819 |       of Const(s,ty) => {Name = s, Ty = ty}
 | 
|  |    820 |        | _ => raise ERR "dest_const" "not a constant");
 | 
|  |    821 | 
 | 
|  |    822 | fun dest_var ctm =
 | 
|  |    823 |    (case Thm.term_of ctm
 | 
| 60521 |    824 |       of Var((s,_),ty) => {Name=s, Ty=ty}
 | 
| 60520 |    825 |        | Free(s,ty)    => {Name=s, Ty=ty}
 | 
|  |    826 |        |             _ => raise ERR "dest_var" "not a variable");
 | 
|  |    827 | 
 | 
|  |    828 | 
 | 
|  |    829 | (*---------------------------------------------------------------------------
 | 
|  |    830 |  * Derived destructor operations.
 | 
|  |    831 |  *---------------------------------------------------------------------------*)
 | 
|  |    832 | 
 | 
|  |    833 | fun dest_monop expected tm =
 | 
|  |    834 |  let
 | 
|  |    835 |    fun err () = raise ERR "dest_monop" ("Not a(n) " ^ quote expected);
 | 
|  |    836 |    val (c, N) = dest_comb tm handle Utils.ERR _ => err ();
 | 
|  |    837 |    val name = #Name (dest_const c handle Utils.ERR _ => err ());
 | 
|  |    838 |  in if name = expected then N else err () end;
 | 
|  |    839 | 
 | 
|  |    840 | fun dest_binop expected tm =
 | 
|  |    841 |  let
 | 
|  |    842 |    fun err () = raise ERR "dest_binop" ("Not a(n) " ^ quote expected);
 | 
|  |    843 |    val (M, N) = dest_comb tm handle Utils.ERR _ => err ()
 | 
|  |    844 |  in (dest_monop expected M, N) handle Utils.ERR _ => err () end;
 | 
|  |    845 | 
 | 
|  |    846 | fun dest_binder expected tm =
 | 
|  |    847 |   dest_abs NONE (dest_monop expected tm)
 | 
|  |    848 |   handle Utils.ERR _ => raise ERR "dest_binder" ("Not a(n) " ^ quote expected);
 | 
|  |    849 | 
 | 
|  |    850 | 
 | 
|  |    851 | val dest_neg    = dest_monop @{const_name Not}
 | 
|  |    852 | val dest_pair   = dest_binop @{const_name Pair}
 | 
|  |    853 | val dest_eq     = dest_binop @{const_name HOL.eq}
 | 
|  |    854 | val dest_imp    = dest_binop @{const_name HOL.implies}
 | 
|  |    855 | val dest_conj   = dest_binop @{const_name HOL.conj}
 | 
|  |    856 | val dest_disj   = dest_binop @{const_name HOL.disj}
 | 
|  |    857 | val dest_exists = dest_binder @{const_name Ex}
 | 
|  |    858 | val dest_forall = dest_binder @{const_name All}
 | 
|  |    859 | 
 | 
|  |    860 | (* Query routines *)
 | 
|  |    861 | 
 | 
|  |    862 | val is_eq     = can dest_eq
 | 
|  |    863 | val is_imp    = can dest_imp
 | 
|  |    864 | val is_forall = can dest_forall
 | 
|  |    865 | val is_exists = can dest_exists
 | 
|  |    866 | val is_neg    = can dest_neg
 | 
|  |    867 | val is_conj   = can dest_conj
 | 
|  |    868 | val is_disj   = can dest_disj
 | 
|  |    869 | val is_pair   = can dest_pair
 | 
|  |    870 | 
 | 
|  |    871 | 
 | 
|  |    872 | (*---------------------------------------------------------------------------
 | 
|  |    873 |  * Iterated creation.
 | 
|  |    874 |  *---------------------------------------------------------------------------*)
 | 
|  |    875 | val list_mk_disj = Utils.end_itlist (fn d1 => fn tm => mk_disj (d1, tm));
 | 
|  |    876 | 
 | 
|  |    877 | (*---------------------------------------------------------------------------
 | 
|  |    878 |  * Iterated destruction. (To the "right" in a term.)
 | 
|  |    879 |  *---------------------------------------------------------------------------*)
 | 
|  |    880 | fun strip break tm =
 | 
|  |    881 |   let fun dest (p as (ctm,accum)) =
 | 
|  |    882 |         let val (M,N) = break ctm
 | 
|  |    883 |         in dest (N, M::accum)
 | 
|  |    884 |         end handle Utils.ERR _ => p
 | 
|  |    885 |   in dest (tm,[])
 | 
|  |    886 |   end;
 | 
|  |    887 | 
 | 
|  |    888 | fun rev2swap (x,l) = (rev l, x);
 | 
|  |    889 | 
 | 
|  |    890 | val strip_comb   = strip (Library.swap o dest_comb)  (* Goes to the "left" *)
 | 
|  |    891 | val strip_imp    = rev2swap o strip dest_imp
 | 
|  |    892 | val strip_abs    = rev2swap o strip (dest_abs NONE)
 | 
|  |    893 | val strip_forall = rev2swap o strip dest_forall
 | 
|  |    894 | val strip_exists = rev2swap o strip dest_exists
 | 
|  |    895 | 
 | 
|  |    896 | val strip_disj   = rev o (op::) o strip dest_disj
 | 
|  |    897 | 
 | 
|  |    898 | 
 | 
|  |    899 | (*---------------------------------------------------------------------------
 | 
|  |    900 |  * Going into and out of prop
 | 
|  |    901 |  *---------------------------------------------------------------------------*)
 | 
|  |    902 | 
 | 
|  |    903 | fun is_Trueprop ct =
 | 
|  |    904 |   (case Thm.term_of ct of
 | 
|  |    905 |     Const (@{const_name Trueprop}, _) $ _ => true
 | 
|  |    906 |   | _ => false);
 | 
|  |    907 | 
 | 
|  |    908 | fun mk_prop ct = if is_Trueprop ct then ct else Thm.apply @{cterm Trueprop} ct;
 | 
|  |    909 | fun drop_prop ct = if is_Trueprop ct then Thm.dest_arg ct else ct;
 | 
|  |    910 | 
 | 
|  |    911 | end;
 | 
|  |    912 | 
 | 
|  |    913 | 
 | 
| 60521 |    914 | 
 | 
| 60520 |    915 | (*** emulation of HOL inference rules for TFL ***)
 | 
|  |    916 | 
 | 
|  |    917 | structure Rules: RULES =
 | 
|  |    918 | struct
 | 
|  |    919 | 
 | 
|  |    920 | fun RULES_ERR func mesg = Utils.ERR {module = "Rules", func = func, mesg = mesg};
 | 
|  |    921 | 
 | 
|  |    922 | 
 | 
|  |    923 | fun cconcl thm = Dcterm.drop_prop (#prop (Thm.crep_thm thm));
 | 
|  |    924 | fun chyps thm = map Dcterm.drop_prop (#hyps (Thm.crep_thm thm));
 | 
|  |    925 | 
 | 
|  |    926 | fun dest_thm thm =
 | 
|  |    927 |   let val {prop,hyps,...} = Thm.rep_thm thm
 | 
|  |    928 |   in (map HOLogic.dest_Trueprop hyps, HOLogic.dest_Trueprop prop) end
 | 
|  |    929 |   handle TERM _ => raise RULES_ERR "dest_thm" "missing Trueprop";
 | 
|  |    930 | 
 | 
|  |    931 | 
 | 
|  |    932 | (* Inference rules *)
 | 
|  |    933 | 
 | 
|  |    934 | (*---------------------------------------------------------------------------
 | 
|  |    935 |  *        Equality (one step)
 | 
|  |    936 |  *---------------------------------------------------------------------------*)
 | 
|  |    937 | 
 | 
|  |    938 | fun REFL tm = Thm.reflexive tm RS meta_eq_to_obj_eq
 | 
|  |    939 |   handle THM (msg, _, _) => raise RULES_ERR "REFL" msg;
 | 
|  |    940 | 
 | 
|  |    941 | fun SYM thm = thm RS sym
 | 
|  |    942 |   handle THM (msg, _, _) => raise RULES_ERR "SYM" msg;
 | 
|  |    943 | 
 | 
|  |    944 | fun ALPHA thm ctm1 =
 | 
|  |    945 |   let
 | 
|  |    946 |     val ctm2 = Thm.cprop_of thm;
 | 
|  |    947 |     val ctm2_eq = Thm.reflexive ctm2;
 | 
|  |    948 |     val ctm1_eq = Thm.reflexive ctm1;
 | 
|  |    949 |   in Thm.equal_elim (Thm.transitive ctm2_eq ctm1_eq) thm end
 | 
|  |    950 |   handle THM (msg, _, _) => raise RULES_ERR "ALPHA" msg;
 | 
|  |    951 | 
 | 
|  |    952 | fun rbeta th =
 | 
|  |    953 |   (case Dcterm.strip_comb (cconcl th) of
 | 
| 60521 |    954 |     (_, [_, r]) => Thm.transitive th (Thm.beta_conversion false r)
 | 
| 60520 |    955 |   | _ => raise RULES_ERR "rbeta" "");
 | 
|  |    956 | 
 | 
|  |    957 | 
 | 
|  |    958 | (*----------------------------------------------------------------------------
 | 
|  |    959 |  *        Implication and the assumption list
 | 
|  |    960 |  *
 | 
|  |    961 |  * Assumptions get stuck on the meta-language assumption list. Implications
 | 
|  |    962 |  * are in the object language, so discharging an assumption "A" from theorem
 | 
|  |    963 |  * "B" results in something that looks like "A --> B".
 | 
|  |    964 |  *---------------------------------------------------------------------------*)
 | 
|  |    965 | 
 | 
|  |    966 | fun ASSUME ctm = Thm.assume (Dcterm.mk_prop ctm);
 | 
|  |    967 | 
 | 
|  |    968 | 
 | 
|  |    969 | (*---------------------------------------------------------------------------
 | 
|  |    970 |  * Implication in TFL is -->. Meta-language implication (==>) is only used
 | 
|  |    971 |  * in the implementation of some of the inference rules below.
 | 
|  |    972 |  *---------------------------------------------------------------------------*)
 | 
|  |    973 | fun MP th1 th2 = th2 RS (th1 RS mp)
 | 
|  |    974 |   handle THM (msg, _, _) => raise RULES_ERR "MP" msg;
 | 
|  |    975 | 
 | 
|  |    976 | (*forces the first argument to be a proposition if necessary*)
 | 
|  |    977 | fun DISCH tm thm = Thm.implies_intr (Dcterm.mk_prop tm) thm COMP impI
 | 
|  |    978 |   handle THM (msg, _, _) => raise RULES_ERR "DISCH" msg;
 | 
|  |    979 | 
 | 
|  |    980 | fun DISCH_ALL thm = fold_rev DISCH (#hyps (Thm.crep_thm thm)) thm;
 | 
|  |    981 | 
 | 
|  |    982 | 
 | 
|  |    983 | fun FILTER_DISCH_ALL P thm =
 | 
|  |    984 |  let fun check tm = P (Thm.term_of tm)
 | 
|  |    985 |  in  fold_rev (fn tm => fn th => if check tm then DISCH tm th else th) (chyps thm) thm
 | 
|  |    986 |  end;
 | 
|  |    987 | 
 | 
|  |    988 | fun UNDISCH thm =
 | 
|  |    989 |    let val tm = Dcterm.mk_prop (#1 (Dcterm.dest_imp (cconcl thm)))
 | 
|  |    990 |    in Thm.implies_elim (thm RS mp) (ASSUME tm) end
 | 
|  |    991 |    handle Utils.ERR _ => raise RULES_ERR "UNDISCH" ""
 | 
|  |    992 |      | THM _ => raise RULES_ERR "UNDISCH" "";
 | 
|  |    993 | 
 | 
|  |    994 | fun PROVE_HYP ath bth = MP (DISCH (cconcl ath) bth) ath;
 | 
|  |    995 | 
 | 
| 60522 |    996 | fun IMP_TRANS th1 th2 = th2 RS (th1 RS @{thm tfl_imp_trans})
 | 
| 60520 |    997 |   handle THM (msg, _, _) => raise RULES_ERR "IMP_TRANS" msg;
 | 
|  |    998 | 
 | 
|  |    999 | 
 | 
|  |   1000 | (*----------------------------------------------------------------------------
 | 
|  |   1001 |  *        Conjunction
 | 
|  |   1002 |  *---------------------------------------------------------------------------*)
 | 
|  |   1003 | 
 | 
|  |   1004 | fun CONJUNCT1 thm = thm RS conjunct1
 | 
|  |   1005 |   handle THM (msg, _, _) => raise RULES_ERR "CONJUNCT1" msg;
 | 
|  |   1006 | 
 | 
|  |   1007 | fun CONJUNCT2 thm = thm RS conjunct2
 | 
|  |   1008 |   handle THM (msg, _, _) => raise RULES_ERR "CONJUNCT2" msg;
 | 
|  |   1009 | 
 | 
|  |   1010 | fun CONJUNCTS th = CONJUNCTS (CONJUNCT1 th) @ CONJUNCTS (CONJUNCT2 th) handle Utils.ERR _ => [th];
 | 
|  |   1011 | 
 | 
|  |   1012 | fun LIST_CONJ [] = raise RULES_ERR "LIST_CONJ" "empty list"
 | 
|  |   1013 |   | LIST_CONJ [th] = th
 | 
|  |   1014 |   | LIST_CONJ (th :: rst) = MP (MP (conjI COMP (impI RS impI)) th) (LIST_CONJ rst)
 | 
|  |   1015 |       handle THM (msg, _, _) => raise RULES_ERR "LIST_CONJ" msg;
 | 
|  |   1016 | 
 | 
|  |   1017 | 
 | 
|  |   1018 | (*----------------------------------------------------------------------------
 | 
|  |   1019 |  *        Disjunction
 | 
|  |   1020 |  *---------------------------------------------------------------------------*)
 | 
|  |   1021 | local
 | 
|  |   1022 |   val prop = Thm.prop_of disjI1
 | 
| 60524 |   1023 |   val [_,Q] = Misc_Legacy.term_vars prop
 | 
| 60520 |   1024 |   val disj1 = Thm.forall_intr (Thm.cterm_of @{context} Q) disjI1
 | 
|  |   1025 | in
 | 
|  |   1026 | fun DISJ1 thm tm = thm RS (Thm.forall_elim (Dcterm.drop_prop tm) disj1)
 | 
|  |   1027 |   handle THM (msg, _, _) => raise RULES_ERR "DISJ1" msg;
 | 
|  |   1028 | end;
 | 
|  |   1029 | 
 | 
|  |   1030 | local
 | 
|  |   1031 |   val prop = Thm.prop_of disjI2
 | 
| 60521 |   1032 |   val [P,_] = Misc_Legacy.term_vars prop
 | 
| 60520 |   1033 |   val disj2 = Thm.forall_intr (Thm.cterm_of @{context} P) disjI2
 | 
|  |   1034 | in
 | 
|  |   1035 | fun DISJ2 tm thm = thm RS (Thm.forall_elim (Dcterm.drop_prop tm) disj2)
 | 
|  |   1036 |   handle THM (msg, _, _) => raise RULES_ERR "DISJ2" msg;
 | 
|  |   1037 | end;
 | 
|  |   1038 | 
 | 
|  |   1039 | 
 | 
|  |   1040 | (*----------------------------------------------------------------------------
 | 
|  |   1041 |  *
 | 
|  |   1042 |  *                   A1 |- M1, ..., An |- Mn
 | 
|  |   1043 |  *     ---------------------------------------------------
 | 
|  |   1044 |  *     [A1 |- M1 \/ ... \/ Mn, ..., An |- M1 \/ ... \/ Mn]
 | 
|  |   1045 |  *
 | 
|  |   1046 |  *---------------------------------------------------------------------------*)
 | 
|  |   1047 | 
 | 
|  |   1048 | 
 | 
|  |   1049 | fun EVEN_ORS thms =
 | 
|  |   1050 |   let fun blue ldisjs [] _ = []
 | 
|  |   1051 |         | blue ldisjs (th::rst) rdisjs =
 | 
|  |   1052 |             let val tail = tl rdisjs
 | 
|  |   1053 |                 val rdisj_tl = Dcterm.list_mk_disj tail
 | 
|  |   1054 |             in fold_rev DISJ2 ldisjs (DISJ1 th rdisj_tl)
 | 
|  |   1055 |                :: blue (ldisjs @ [cconcl th]) rst tail
 | 
|  |   1056 |             end handle Utils.ERR _ => [fold_rev DISJ2 ldisjs th]
 | 
|  |   1057 |    in blue [] thms (map cconcl thms) end;
 | 
|  |   1058 | 
 | 
|  |   1059 | 
 | 
|  |   1060 | (*----------------------------------------------------------------------------
 | 
|  |   1061 |  *
 | 
|  |   1062 |  *         A |- P \/ Q   B,P |- R    C,Q |- R
 | 
|  |   1063 |  *     ---------------------------------------------------
 | 
|  |   1064 |  *                     A U B U C |- R
 | 
|  |   1065 |  *
 | 
|  |   1066 |  *---------------------------------------------------------------------------*)
 | 
|  |   1067 | 
 | 
|  |   1068 | fun DISJ_CASES th1 th2 th3 =
 | 
|  |   1069 |   let
 | 
|  |   1070 |     val c = Dcterm.drop_prop (cconcl th1);
 | 
|  |   1071 |     val (disj1, disj2) = Dcterm.dest_disj c;
 | 
|  |   1072 |     val th2' = DISCH disj1 th2;
 | 
|  |   1073 |     val th3' = DISCH disj2 th3;
 | 
|  |   1074 |   in
 | 
| 60522 |   1075 |     th3' RS (th2' RS (th1 RS @{thm tfl_disjE}))
 | 
| 60520 |   1076 |       handle THM (msg, _, _) => raise RULES_ERR "DISJ_CASES" msg
 | 
|  |   1077 |   end;
 | 
|  |   1078 | 
 | 
|  |   1079 | 
 | 
|  |   1080 | (*-----------------------------------------------------------------------------
 | 
|  |   1081 |  *
 | 
|  |   1082 |  *       |- A1 \/ ... \/ An     [A1 |- M, ..., An |- M]
 | 
|  |   1083 |  *     ---------------------------------------------------
 | 
|  |   1084 |  *                           |- M
 | 
|  |   1085 |  *
 | 
|  |   1086 |  * Note. The list of theorems may be all jumbled up, so we have to
 | 
|  |   1087 |  * first organize it to align with the first argument (the disjunctive
 | 
|  |   1088 |  * theorem).
 | 
|  |   1089 |  *---------------------------------------------------------------------------*)
 | 
|  |   1090 | 
 | 
|  |   1091 | fun organize eq =    (* a bit slow - analogous to insertion sort *)
 | 
|  |   1092 |  let fun extract a alist =
 | 
|  |   1093 |      let fun ex (_,[]) = raise RULES_ERR "organize" "not a permutation.1"
 | 
|  |   1094 |            | ex(left,h::t) = if (eq h a) then (h,rev left@t) else ex(h::left,t)
 | 
|  |   1095 |      in ex ([],alist)
 | 
|  |   1096 |      end
 | 
|  |   1097 |      fun place [] [] = []
 | 
|  |   1098 |        | place (a::rst) alist =
 | 
|  |   1099 |            let val (item,next) = extract a alist
 | 
|  |   1100 |            in item::place rst next
 | 
|  |   1101 |            end
 | 
|  |   1102 |        | place _ _ = raise RULES_ERR "organize" "not a permutation.2"
 | 
|  |   1103 |  in place
 | 
|  |   1104 |  end;
 | 
|  |   1105 | 
 | 
|  |   1106 | fun DISJ_CASESL disjth thl =
 | 
|  |   1107 |    let val c = cconcl disjth
 | 
|  |   1108 |        fun eq th atm =
 | 
|  |   1109 |         exists (fn t => HOLogic.dest_Trueprop t aconv Thm.term_of atm) (Thm.hyps_of th)
 | 
|  |   1110 |        val tml = Dcterm.strip_disj c
 | 
| 60521 |   1111 |        fun DL _ [] = raise RULES_ERR "DISJ_CASESL" "no cases"
 | 
| 60520 |   1112 |          | DL th [th1] = PROVE_HYP th th1
 | 
|  |   1113 |          | DL th [th1,th2] = DISJ_CASES th th1 th2
 | 
|  |   1114 |          | DL th (th1::rst) =
 | 
|  |   1115 |             let val tm = #2 (Dcterm.dest_disj (Dcterm.drop_prop(cconcl th)))
 | 
|  |   1116 |              in DISJ_CASES th th1 (DL (ASSUME tm) rst) end
 | 
|  |   1117 |    in DL disjth (organize eq tml thl)
 | 
|  |   1118 |    end;
 | 
|  |   1119 | 
 | 
|  |   1120 | 
 | 
|  |   1121 | (*----------------------------------------------------------------------------
 | 
|  |   1122 |  *        Universals
 | 
|  |   1123 |  *---------------------------------------------------------------------------*)
 | 
|  |   1124 | local (* this is fragile *)
 | 
|  |   1125 |   val prop = Thm.prop_of spec
 | 
|  |   1126 |   val x = hd (tl (Misc_Legacy.term_vars prop))
 | 
|  |   1127 |   val cTV = Thm.ctyp_of @{context} (type_of x)
 | 
|  |   1128 |   val gspec = Thm.forall_intr (Thm.cterm_of @{context} x) spec
 | 
|  |   1129 | in
 | 
|  |   1130 | fun SPEC tm thm =
 | 
|  |   1131 |    let val gspec' = Drule.instantiate_normalize ([(cTV, Thm.ctyp_of_cterm tm)], []) gspec
 | 
|  |   1132 |    in thm RS (Thm.forall_elim tm gspec') end
 | 
|  |   1133 | end;
 | 
|  |   1134 | 
 | 
|  |   1135 | fun SPEC_ALL thm = fold SPEC (#1 (Dcterm.strip_forall(cconcl thm))) thm;
 | 
|  |   1136 | 
 | 
|  |   1137 | val ISPEC = SPEC
 | 
|  |   1138 | val ISPECL = fold ISPEC;
 | 
|  |   1139 | 
 | 
|  |   1140 | (* Not optimized! Too complicated. *)
 | 
|  |   1141 | local
 | 
|  |   1142 |   val prop = Thm.prop_of allI
 | 
|  |   1143 |   val [P] = Misc_Legacy.add_term_vars (prop, [])
 | 
|  |   1144 |   fun cty_theta ctxt = map (fn (i, (S, ty)) => apply2 (Thm.ctyp_of ctxt) (TVar (i, S), ty))
 | 
|  |   1145 |   fun ctm_theta ctxt =
 | 
|  |   1146 |     map (fn (i, (_, tm2)) =>
 | 
|  |   1147 |       let val ctm2 = Thm.cterm_of ctxt tm2
 | 
|  |   1148 |       in (Thm.cterm_of ctxt (Var (i, Thm.typ_of_cterm ctm2)), ctm2) end)
 | 
|  |   1149 |   fun certify ctxt (ty_theta,tm_theta) =
 | 
|  |   1150 |     (cty_theta ctxt (Vartab.dest ty_theta),
 | 
|  |   1151 |      ctm_theta ctxt (Vartab.dest tm_theta))
 | 
|  |   1152 | in
 | 
|  |   1153 | fun GEN ctxt v th =
 | 
|  |   1154 |    let val gth = Thm.forall_intr v th
 | 
|  |   1155 |        val thy = Proof_Context.theory_of ctxt
 | 
|  |   1156 |        val Const(@{const_name Pure.all},_)$Abs(x,ty,rst) = Thm.prop_of gth
 | 
|  |   1157 |        val P' = Abs(x,ty, HOLogic.dest_Trueprop rst)  (* get rid of trueprop *)
 | 
|  |   1158 |        val theta = Pattern.match thy (P,P') (Vartab.empty, Vartab.empty);
 | 
|  |   1159 |        val allI2 = Drule.instantiate_normalize (certify ctxt theta) allI
 | 
|  |   1160 |        val thm = Thm.implies_elim allI2 gth
 | 
|  |   1161 |        val tp $ (A $ Abs(_,_,M)) = Thm.prop_of thm
 | 
|  |   1162 |        val prop' = tp $ (A $ Abs(x,ty,M))
 | 
|  |   1163 |    in ALPHA thm (Thm.cterm_of ctxt prop') end
 | 
|  |   1164 | end;
 | 
|  |   1165 | 
 | 
|  |   1166 | fun GENL ctxt = fold_rev (GEN ctxt);
 | 
|  |   1167 | 
 | 
|  |   1168 | fun GEN_ALL ctxt thm =
 | 
|  |   1169 |   let
 | 
|  |   1170 |     val prop = Thm.prop_of thm
 | 
|  |   1171 |     val vlist = map (Thm.cterm_of ctxt) (Misc_Legacy.add_term_vars (prop, []))
 | 
|  |   1172 |   in GENL ctxt vlist thm end;
 | 
|  |   1173 | 
 | 
|  |   1174 | 
 | 
|  |   1175 | fun MATCH_MP th1 th2 =
 | 
|  |   1176 |    if (Dcterm.is_forall (Dcterm.drop_prop(cconcl th1)))
 | 
|  |   1177 |    then MATCH_MP (th1 RS spec) th2
 | 
|  |   1178 |    else MP th1 th2;
 | 
|  |   1179 | 
 | 
|  |   1180 | 
 | 
|  |   1181 | (*----------------------------------------------------------------------------
 | 
|  |   1182 |  *        Existentials
 | 
|  |   1183 |  *---------------------------------------------------------------------------*)
 | 
|  |   1184 | 
 | 
|  |   1185 | 
 | 
|  |   1186 | 
 | 
|  |   1187 | (*---------------------------------------------------------------------------
 | 
|  |   1188 |  * Existential elimination
 | 
|  |   1189 |  *
 | 
|  |   1190 |  *      A1 |- ?x.t[x]   ,   A2, "t[v]" |- t'
 | 
|  |   1191 |  *      ------------------------------------     (variable v occurs nowhere)
 | 
|  |   1192 |  *                A1 u A2 |- t'
 | 
|  |   1193 |  *
 | 
|  |   1194 |  *---------------------------------------------------------------------------*)
 | 
|  |   1195 | 
 | 
|  |   1196 | fun CHOOSE ctxt (fvar, exth) fact =
 | 
|  |   1197 |   let
 | 
|  |   1198 |     val lam = #2 (Dcterm.dest_comb (Dcterm.drop_prop (cconcl exth)))
 | 
|  |   1199 |     val redex = Dcterm.capply lam fvar
 | 
|  |   1200 |     val t$u = Thm.term_of redex
 | 
|  |   1201 |     val residue = Thm.cterm_of ctxt (Term.betapply (t, u))
 | 
|  |   1202 |   in
 | 
| 60522 |   1203 |     GEN ctxt fvar (DISCH residue fact) RS (exth RS @{thm tfl_exE})
 | 
| 60520 |   1204 |       handle THM (msg, _, _) => raise RULES_ERR "CHOOSE" msg
 | 
|  |   1205 |   end;
 | 
|  |   1206 | 
 | 
|  |   1207 | local
 | 
|  |   1208 |   val prop = Thm.prop_of exI
 | 
|  |   1209 |   val [P, x] = map (Thm.cterm_of @{context}) (Misc_Legacy.term_vars prop)
 | 
|  |   1210 | in
 | 
|  |   1211 | fun EXISTS (template,witness) thm =
 | 
|  |   1212 |   let val abstr = #2 (Dcterm.dest_comb template) in
 | 
|  |   1213 |     thm RS (cterm_instantiate [(P, abstr), (x, witness)] exI)
 | 
|  |   1214 |       handle THM (msg, _, _) => raise RULES_ERR "EXISTS" msg
 | 
|  |   1215 |   end
 | 
|  |   1216 | end;
 | 
|  |   1217 | 
 | 
|  |   1218 | (*----------------------------------------------------------------------------
 | 
|  |   1219 |  *
 | 
|  |   1220 |  *         A |- M
 | 
|  |   1221 |  *   -------------------   [v_1,...,v_n]
 | 
|  |   1222 |  *    A |- ?v1...v_n. M
 | 
|  |   1223 |  *
 | 
|  |   1224 |  *---------------------------------------------------------------------------*)
 | 
|  |   1225 | 
 | 
|  |   1226 | fun EXISTL vlist th =
 | 
|  |   1227 |   fold_rev (fn v => fn thm => EXISTS(Dcterm.mk_exists(v,cconcl thm), v) thm)
 | 
|  |   1228 |            vlist th;
 | 
|  |   1229 | 
 | 
|  |   1230 | 
 | 
|  |   1231 | (*----------------------------------------------------------------------------
 | 
|  |   1232 |  *
 | 
|  |   1233 |  *       A |- M[x_1,...,x_n]
 | 
|  |   1234 |  *   ----------------------------   [(x |-> y)_1,...,(x |-> y)_n]
 | 
|  |   1235 |  *       A |- ?y_1...y_n. M
 | 
|  |   1236 |  *
 | 
|  |   1237 |  *---------------------------------------------------------------------------*)
 | 
|  |   1238 | (* Could be improved, but needs "subst_free" for certified terms *)
 | 
|  |   1239 | 
 | 
|  |   1240 | fun IT_EXISTS ctxt blist th =
 | 
|  |   1241 |   let
 | 
|  |   1242 |     val blist' = map (apply2 Thm.term_of) blist
 | 
|  |   1243 |     fun ex v M = Thm.cterm_of ctxt (USyntax.mk_exists{Bvar=v,Body = M})
 | 
|  |   1244 |   in
 | 
|  |   1245 |     fold_rev (fn (b as (r1,r2)) => fn thm =>
 | 
|  |   1246 |         EXISTS(ex r2 (subst_free [b]
 | 
|  |   1247 |                    (HOLogic.dest_Trueprop(Thm.prop_of thm))), Thm.cterm_of ctxt r1)
 | 
|  |   1248 |               thm)
 | 
|  |   1249 |        blist' th
 | 
|  |   1250 |   end;
 | 
|  |   1251 | 
 | 
|  |   1252 | (*---------------------------------------------------------------------------
 | 
|  |   1253 |  *  Faster version, that fails for some as yet unknown reason
 | 
|  |   1254 |  * fun IT_EXISTS blist th =
 | 
|  |   1255 |  *    let val {thy,...} = rep_thm th
 | 
|  |   1256 |  *        val tych = cterm_of thy
 | 
|  |   1257 |  *        fun detype (x,y) = ((#t o rep_cterm) x, (#t o rep_cterm) y)
 | 
|  |   1258 |  *   in
 | 
|  |   1259 |  *  fold (fn (b as (r1,r2), thm) =>
 | 
|  |   1260 |  *  EXISTS(D.mk_exists(r2, tych(subst_free[detype b](#t(rep_cterm(cconcl thm))))),
 | 
|  |   1261 |  *           r1) thm)  blist th
 | 
|  |   1262 |  *   end;
 | 
|  |   1263 |  *---------------------------------------------------------------------------*)
 | 
|  |   1264 | 
 | 
|  |   1265 | (*----------------------------------------------------------------------------
 | 
|  |   1266 |  *        Rewriting
 | 
|  |   1267 |  *---------------------------------------------------------------------------*)
 | 
|  |   1268 | 
 | 
|  |   1269 | fun SUBS ctxt thl =
 | 
|  |   1270 |   rewrite_rule ctxt (map (fn th => th RS eq_reflection handle THM _ => th) thl);
 | 
|  |   1271 | 
 | 
|  |   1272 | val rew_conv = Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE));
 | 
|  |   1273 | 
 | 
|  |   1274 | fun simpl_conv ctxt thl ctm =
 | 
|  |   1275 |  rew_conv (ctxt addsimps thl) ctm RS meta_eq_to_obj_eq;
 | 
|  |   1276 | 
 | 
|  |   1277 | 
 | 
| 60522 |   1278 | fun RIGHT_ASSOC ctxt = rewrite_rule ctxt @{thms tfl_disj_assoc};
 | 
| 60520 |   1279 | 
 | 
|  |   1280 | 
 | 
|  |   1281 | 
 | 
|  |   1282 | (*---------------------------------------------------------------------------
 | 
|  |   1283 |  *                  TERMINATION CONDITION EXTRACTION
 | 
|  |   1284 |  *---------------------------------------------------------------------------*)
 | 
|  |   1285 | 
 | 
|  |   1286 | 
 | 
|  |   1287 | (* Object language quantifier, i.e., "!" *)
 | 
|  |   1288 | fun Forall v M = USyntax.mk_forall{Bvar=v, Body=M};
 | 
|  |   1289 | 
 | 
|  |   1290 | 
 | 
|  |   1291 | (* Fragile: it's a cong if it is not "R y x ==> cut f R x y = f y" *)
 | 
|  |   1292 | fun is_cong thm =
 | 
|  |   1293 |   case (Thm.prop_of thm) of
 | 
|  |   1294 |     (Const(@{const_name Pure.imp},_)$(Const(@{const_name Trueprop},_)$ _) $
 | 
| 60524 |   1295 |       (Const(@{const_name Pure.eq},_) $ (Const (@{const_name Wfrec.cut},_) $ _ $ _ $ _ $ _) $ _)) =>
 | 
| 60520 |   1296 |         false
 | 
|  |   1297 |   | _ => true;
 | 
|  |   1298 | 
 | 
|  |   1299 | 
 | 
|  |   1300 | fun dest_equal(Const (@{const_name Pure.eq},_) $
 | 
|  |   1301 |                (Const (@{const_name Trueprop},_) $ lhs)
 | 
|  |   1302 |                $ (Const (@{const_name Trueprop},_) $ rhs)) = {lhs=lhs, rhs=rhs}
 | 
|  |   1303 |   | dest_equal(Const (@{const_name Pure.eq},_) $ lhs $ rhs) = {lhs=lhs, rhs=rhs}
 | 
|  |   1304 |   | dest_equal tm = USyntax.dest_eq tm;
 | 
|  |   1305 | 
 | 
|  |   1306 | fun get_lhs tm = #lhs(dest_equal (HOLogic.dest_Trueprop tm));
 | 
|  |   1307 | 
 | 
|  |   1308 | fun dest_all used (Const(@{const_name Pure.all},_) $ (a as Abs _)) = USyntax.dest_abs used a
 | 
|  |   1309 |   | dest_all _ _ = raise RULES_ERR "dest_all" "not a !!";
 | 
|  |   1310 | 
 | 
|  |   1311 | val is_all = can (dest_all []);
 | 
|  |   1312 | 
 | 
|  |   1313 | fun strip_all used fm =
 | 
|  |   1314 |    if (is_all fm)
 | 
|  |   1315 |    then let val ({Bvar, Body}, used') = dest_all used fm
 | 
|  |   1316 |             val (bvs, core, used'') = strip_all used' Body
 | 
|  |   1317 |         in ((Bvar::bvs), core, used'')
 | 
|  |   1318 |         end
 | 
|  |   1319 |    else ([], fm, used);
 | 
|  |   1320 | 
 | 
|  |   1321 | fun list_break_all(Const(@{const_name Pure.all},_) $ Abs (s,ty,body)) =
 | 
|  |   1322 |      let val (L,core) = list_break_all body
 | 
|  |   1323 |      in ((s,ty)::L, core)
 | 
|  |   1324 |      end
 | 
|  |   1325 |   | list_break_all tm = ([],tm);
 | 
|  |   1326 | 
 | 
|  |   1327 | (*---------------------------------------------------------------------------
 | 
|  |   1328 |  * Rename a term of the form
 | 
|  |   1329 |  *
 | 
|  |   1330 |  *      !!x1 ...xn. x1=M1 ==> ... ==> xn=Mn
 | 
|  |   1331 |  *                  ==> ((%v1...vn. Q) x1 ... xn = g x1 ... xn.
 | 
|  |   1332 |  * to one of
 | 
|  |   1333 |  *
 | 
|  |   1334 |  *      !!v1 ... vn. v1=M1 ==> ... ==> vn=Mn
 | 
|  |   1335 |  *      ==> ((%v1...vn. Q) v1 ... vn = g v1 ... vn.
 | 
|  |   1336 |  *
 | 
|  |   1337 |  * This prevents name problems in extraction, and helps the result to read
 | 
|  |   1338 |  * better. There is a problem with varstructs, since they can introduce more
 | 
|  |   1339 |  * than n variables, and some extra reasoning needs to be done.
 | 
|  |   1340 |  *---------------------------------------------------------------------------*)
 | 
|  |   1341 | 
 | 
|  |   1342 | fun get ([],_,L) = rev L
 | 
|  |   1343 |   | get (ant::rst,n,L) =
 | 
|  |   1344 |       case (list_break_all ant)
 | 
|  |   1345 |         of ([],_) => get (rst, n+1,L)
 | 
| 60521 |   1346 |          | (_,body) =>
 | 
| 60520 |   1347 |             let val eq = Logic.strip_imp_concl body
 | 
| 60521 |   1348 |                 val (f,_) = USyntax.strip_comb (get_lhs eq)
 | 
| 60520 |   1349 |                 val (vstrl,_) = USyntax.strip_abs f
 | 
|  |   1350 |                 val names  =
 | 
|  |   1351 |                   Name.variant_list (Misc_Legacy.add_term_names(body, [])) (map (#1 o dest_Free) vstrl)
 | 
|  |   1352 |             in get (rst, n+1, (names,n)::L) end
 | 
|  |   1353 |             handle TERM _ => get (rst, n+1, L)
 | 
|  |   1354 |               | Utils.ERR _ => get (rst, n+1, L);
 | 
|  |   1355 | 
 | 
|  |   1356 | (* Note: Thm.rename_params_rule counts from 1, not 0 *)
 | 
|  |   1357 | fun rename thm =
 | 
|  |   1358 |   let
 | 
|  |   1359 |     val ants = Logic.strip_imp_prems (Thm.prop_of thm)
 | 
|  |   1360 |     val news = get (ants,1,[])
 | 
|  |   1361 |   in fold Thm.rename_params_rule news thm end;
 | 
|  |   1362 | 
 | 
|  |   1363 | 
 | 
|  |   1364 | (*---------------------------------------------------------------------------
 | 
|  |   1365 |  * Beta-conversion to the rhs of an equation (taken from hol90/drule.sml)
 | 
|  |   1366 |  *---------------------------------------------------------------------------*)
 | 
|  |   1367 | 
 | 
|  |   1368 | fun list_beta_conv tm =
 | 
|  |   1369 |   let fun rbeta th = Thm.transitive th (Thm.beta_conversion false (#2(Dcterm.dest_eq(cconcl th))))
 | 
|  |   1370 |       fun iter [] = Thm.reflexive tm
 | 
|  |   1371 |         | iter (v::rst) = rbeta (Thm.combination(iter rst) (Thm.reflexive v))
 | 
|  |   1372 |   in iter  end;
 | 
|  |   1373 | 
 | 
|  |   1374 | 
 | 
|  |   1375 | (*---------------------------------------------------------------------------
 | 
|  |   1376 |  * Trace information for the rewriter
 | 
|  |   1377 |  *---------------------------------------------------------------------------*)
 | 
|  |   1378 | val tracing = Unsynchronized.ref false;
 | 
|  |   1379 | 
 | 
|  |   1380 | fun say s = if !tracing then writeln s else ();
 | 
|  |   1381 | 
 | 
|  |   1382 | fun print_thms ctxt s L =
 | 
|  |   1383 |   say (cat_lines (s :: map (Display.string_of_thm ctxt) L));
 | 
|  |   1384 | 
 | 
|  |   1385 | fun print_term ctxt s t =
 | 
|  |   1386 |   say (cat_lines [s, Syntax.string_of_term ctxt t]);
 | 
|  |   1387 | 
 | 
|  |   1388 | 
 | 
|  |   1389 | (*---------------------------------------------------------------------------
 | 
|  |   1390 |  * General abstraction handlers, should probably go in USyntax.
 | 
|  |   1391 |  *---------------------------------------------------------------------------*)
 | 
|  |   1392 | fun mk_aabs (vstr, body) =
 | 
|  |   1393 |   USyntax.mk_abs {Bvar = vstr, Body = body}
 | 
|  |   1394 |   handle Utils.ERR _ => USyntax.mk_pabs {varstruct = vstr, body = body};
 | 
|  |   1395 | 
 | 
|  |   1396 | fun list_mk_aabs (vstrl,tm) =
 | 
|  |   1397 |     fold_rev (fn vstr => fn tm => mk_aabs(vstr,tm)) vstrl tm;
 | 
|  |   1398 | 
 | 
|  |   1399 | fun dest_aabs used tm =
 | 
|  |   1400 |    let val ({Bvar,Body}, used') = USyntax.dest_abs used tm
 | 
|  |   1401 |    in (Bvar, Body, used') end
 | 
|  |   1402 |    handle Utils.ERR _ =>
 | 
|  |   1403 |      let val {varstruct, body, used} = USyntax.dest_pabs used tm
 | 
|  |   1404 |      in (varstruct, body, used) end;
 | 
|  |   1405 | 
 | 
|  |   1406 | fun strip_aabs used tm =
 | 
|  |   1407 |    let val (vstr, body, used') = dest_aabs used tm
 | 
|  |   1408 |        val (bvs, core, used'') = strip_aabs used' body
 | 
|  |   1409 |    in (vstr::bvs, core, used'') end
 | 
|  |   1410 |    handle Utils.ERR _ => ([], tm, used);
 | 
|  |   1411 | 
 | 
|  |   1412 | fun dest_combn tm 0 = (tm,[])
 | 
|  |   1413 |   | dest_combn tm n =
 | 
|  |   1414 |      let val {Rator,Rand} = USyntax.dest_comb tm
 | 
|  |   1415 |          val (f,rands) = dest_combn Rator (n-1)
 | 
|  |   1416 |      in (f,Rand::rands)
 | 
|  |   1417 |      end;
 | 
|  |   1418 | 
 | 
|  |   1419 | 
 | 
|  |   1420 | 
 | 
|  |   1421 | 
 | 
|  |   1422 | local fun dest_pair M = let val {fst,snd} = USyntax.dest_pair M in (fst,snd) end
 | 
|  |   1423 |       fun mk_fst tm =
 | 
|  |   1424 |           let val ty as Type(@{type_name Product_Type.prod}, [fty,sty]) = type_of tm
 | 
|  |   1425 |           in  Const (@{const_name Product_Type.fst}, ty --> fty) $ tm  end
 | 
|  |   1426 |       fun mk_snd tm =
 | 
|  |   1427 |           let val ty as Type(@{type_name Product_Type.prod}, [fty,sty]) = type_of tm
 | 
|  |   1428 |           in  Const (@{const_name Product_Type.snd}, ty --> sty) $ tm  end
 | 
|  |   1429 | in
 | 
|  |   1430 | fun XFILL tych x vstruct =
 | 
|  |   1431 |   let fun traverse p xocc L =
 | 
|  |   1432 |         if (is_Free p)
 | 
|  |   1433 |         then tych xocc::L
 | 
|  |   1434 |         else let val (p1,p2) = dest_pair p
 | 
|  |   1435 |              in traverse p1 (mk_fst xocc) (traverse p2  (mk_snd xocc) L)
 | 
|  |   1436 |              end
 | 
|  |   1437 |   in
 | 
|  |   1438 |   traverse vstruct x []
 | 
|  |   1439 | end end;
 | 
|  |   1440 | 
 | 
|  |   1441 | (*---------------------------------------------------------------------------
 | 
|  |   1442 |  * Replace a free tuple (vstr) by a universally quantified variable (a).
 | 
|  |   1443 |  * Note that the notion of "freeness" for a tuple is different than for a
 | 
|  |   1444 |  * variable: if variables in the tuple also occur in any other place than
 | 
|  |   1445 |  * an occurrences of the tuple, they aren't "free" (which is thus probably
 | 
|  |   1446 |  *  the wrong word to use).
 | 
|  |   1447 |  *---------------------------------------------------------------------------*)
 | 
|  |   1448 | 
 | 
|  |   1449 | fun VSTRUCT_ELIM ctxt tych a vstr th =
 | 
|  |   1450 |   let val L = USyntax.free_vars_lr vstr
 | 
|  |   1451 |       val bind1 = tych (HOLogic.mk_Trueprop (HOLogic.mk_eq(a,vstr)))
 | 
|  |   1452 |       val thm1 = Thm.implies_intr bind1 (SUBS ctxt [SYM(Thm.assume bind1)] th)
 | 
|  |   1453 |       val thm2 = forall_intr_list (map tych L) thm1
 | 
|  |   1454 |       val thm3 = forall_elim_list (XFILL tych a vstr) thm2
 | 
|  |   1455 |   in refl RS
 | 
|  |   1456 |      rewrite_rule ctxt [Thm.symmetric (@{thm surjective_pairing} RS eq_reflection)] thm3
 | 
|  |   1457 |   end;
 | 
|  |   1458 | 
 | 
|  |   1459 | fun PGEN ctxt tych a vstr th =
 | 
|  |   1460 |   let val a1 = tych a
 | 
|  |   1461 |       val vstr1 = tych vstr
 | 
|  |   1462 |   in
 | 
|  |   1463 |   Thm.forall_intr a1
 | 
|  |   1464 |      (if (is_Free vstr)
 | 
|  |   1465 |       then cterm_instantiate [(vstr1,a1)] th
 | 
|  |   1466 |       else VSTRUCT_ELIM ctxt tych a vstr th)
 | 
|  |   1467 |   end;
 | 
|  |   1468 | 
 | 
|  |   1469 | 
 | 
|  |   1470 | (*---------------------------------------------------------------------------
 | 
|  |   1471 |  * Takes apart a paired beta-redex, looking like "(\(x,y).N) vstr", into
 | 
|  |   1472 |  *
 | 
|  |   1473 |  *     (([x,y],N),vstr)
 | 
|  |   1474 |  *---------------------------------------------------------------------------*)
 | 
|  |   1475 | fun dest_pbeta_redex used M n =
 | 
|  |   1476 |   let val (f,args) = dest_combn M n
 | 
| 60521 |   1477 |       val _ = dest_aabs used f
 | 
| 60520 |   1478 |   in (strip_aabs used f,args)
 | 
|  |   1479 |   end;
 | 
|  |   1480 | 
 | 
|  |   1481 | fun pbeta_redex M n = can (fn t => dest_pbeta_redex [] t n) M;
 | 
|  |   1482 | 
 | 
|  |   1483 | fun dest_impl tm =
 | 
|  |   1484 |   let val ants = Logic.strip_imp_prems tm
 | 
|  |   1485 |       val eq = Logic.strip_imp_concl tm
 | 
|  |   1486 |   in (ants,get_lhs eq)
 | 
|  |   1487 |   end;
 | 
|  |   1488 | 
 | 
|  |   1489 | fun restricted t = is_some (USyntax.find_term
 | 
|  |   1490 |                             (fn (Const(@{const_name Wfrec.cut},_)) =>true | _ => false)
 | 
|  |   1491 |                             t)
 | 
|  |   1492 | 
 | 
|  |   1493 | fun CONTEXT_REWRITE_RULE main_ctxt (func, G, cut_lemma, congs) th =
 | 
|  |   1494 |  let val globals = func::G
 | 
|  |   1495 |      val ctxt0 = empty_simpset main_ctxt
 | 
|  |   1496 |      val pbeta_reduce = simpl_conv ctxt0 [@{thm split_conv} RS eq_reflection];
 | 
|  |   1497 |      val tc_list = Unsynchronized.ref []: term list Unsynchronized.ref
 | 
|  |   1498 |      val cut_lemma' = cut_lemma RS eq_reflection
 | 
|  |   1499 |      fun prover used ctxt thm =
 | 
|  |   1500 |      let fun cong_prover ctxt thm =
 | 
| 60521 |   1501 |          let val _ = say "cong_prover:"
 | 
| 60520 |   1502 |              val cntxt = Simplifier.prems_of ctxt
 | 
| 60521 |   1503 |              val _ = print_thms ctxt "cntxt:" cntxt
 | 
|  |   1504 |              val _ = say "cong rule:"
 | 
|  |   1505 |              val _ = say (Display.string_of_thm ctxt thm)
 | 
| 60520 |   1506 |              (* Unquantified eliminate *)
 | 
|  |   1507 |              fun uq_eliminate (thm,imp) =
 | 
|  |   1508 |                  let val tych = Thm.cterm_of ctxt
 | 
|  |   1509 |                      val _ = print_term ctxt "To eliminate:" imp
 | 
|  |   1510 |                      val ants = map tych (Logic.strip_imp_prems imp)
 | 
|  |   1511 |                      val eq = Logic.strip_imp_concl imp
 | 
|  |   1512 |                      val lhs = tych(get_lhs eq)
 | 
|  |   1513 |                      val ctxt' = Simplifier.add_prems (map ASSUME ants) ctxt
 | 
|  |   1514 |                      val lhs_eq_lhs1 = Raw_Simplifier.rewrite_cterm (false,true,false) (prover used) ctxt' lhs
 | 
|  |   1515 |                        handle Utils.ERR _ => Thm.reflexive lhs
 | 
|  |   1516 |                      val _ = print_thms ctxt' "proven:" [lhs_eq_lhs1]
 | 
|  |   1517 |                      val lhs_eq_lhs2 = implies_intr_list ants lhs_eq_lhs1
 | 
|  |   1518 |                      val lhs_eeq_lhs2 = lhs_eq_lhs2 RS meta_eq_to_obj_eq
 | 
|  |   1519 |                   in
 | 
|  |   1520 |                   lhs_eeq_lhs2 COMP thm
 | 
|  |   1521 |                   end
 | 
|  |   1522 |              fun pq_eliminate (thm, vlist, imp_body, lhs_eq) =
 | 
|  |   1523 |               let val ((vstrl, _, used'), args) = dest_pbeta_redex used lhs_eq (length vlist)
 | 
| 60521 |   1524 |                   val _ = forall (op aconv) (ListPair.zip (vlist, args))
 | 
| 60520 |   1525 |                     orelse error "assertion failed in CONTEXT_REWRITE_RULE"
 | 
|  |   1526 |                   val imp_body1 = subst_free (ListPair.zip (args, vstrl))
 | 
|  |   1527 |                                              imp_body
 | 
|  |   1528 |                   val tych = Thm.cterm_of ctxt
 | 
|  |   1529 |                   val ants1 = map tych (Logic.strip_imp_prems imp_body1)
 | 
|  |   1530 |                   val eq1 = Logic.strip_imp_concl imp_body1
 | 
|  |   1531 |                   val Q = get_lhs eq1
 | 
|  |   1532 |                   val QeqQ1 = pbeta_reduce (tych Q)
 | 
|  |   1533 |                   val Q1 = #2(Dcterm.dest_eq(cconcl QeqQ1))
 | 
|  |   1534 |                   val ctxt' = Simplifier.add_prems (map ASSUME ants1) ctxt
 | 
|  |   1535 |                   val Q1eeqQ2 = Raw_Simplifier.rewrite_cterm (false,true,false) (prover used') ctxt' Q1
 | 
|  |   1536 |                                 handle Utils.ERR _ => Thm.reflexive Q1
 | 
|  |   1537 |                   val Q2 = #2 (Logic.dest_equals (Thm.prop_of Q1eeqQ2))
 | 
|  |   1538 |                   val Q3 = tych(list_comb(list_mk_aabs(vstrl,Q2),vstrl))
 | 
|  |   1539 |                   val Q2eeqQ3 = Thm.symmetric(pbeta_reduce Q3 RS eq_reflection)
 | 
|  |   1540 |                   val thA = Thm.transitive(QeqQ1 RS eq_reflection) Q1eeqQ2
 | 
|  |   1541 |                   val QeeqQ3 = Thm.transitive thA Q2eeqQ3 handle THM _ =>
 | 
|  |   1542 |                                ((Q2eeqQ3 RS meta_eq_to_obj_eq)
 | 
|  |   1543 |                                 RS ((thA RS meta_eq_to_obj_eq) RS trans))
 | 
|  |   1544 |                                 RS eq_reflection
 | 
|  |   1545 |                   val impth = implies_intr_list ants1 QeeqQ3
 | 
|  |   1546 |                   val impth1 = impth RS meta_eq_to_obj_eq
 | 
|  |   1547 |                   (* Need to abstract *)
 | 
|  |   1548 |                   val ant_th = Utils.itlist2 (PGEN ctxt' tych) args vstrl impth1
 | 
|  |   1549 |               in ant_th COMP thm
 | 
|  |   1550 |               end
 | 
|  |   1551 |              fun q_eliminate (thm, imp) =
 | 
|  |   1552 |               let val (vlist, imp_body, used') = strip_all used imp
 | 
|  |   1553 |                   val (ants,Q) = dest_impl imp_body
 | 
|  |   1554 |               in if (pbeta_redex Q) (length vlist)
 | 
|  |   1555 |                  then pq_eliminate (thm, vlist, imp_body, Q)
 | 
|  |   1556 |                  else
 | 
|  |   1557 |                  let val tych = Thm.cterm_of ctxt
 | 
|  |   1558 |                      val ants1 = map tych ants
 | 
|  |   1559 |                      val ctxt' = Simplifier.add_prems (map ASSUME ants1) ctxt
 | 
|  |   1560 |                      val Q_eeq_Q1 = Raw_Simplifier.rewrite_cterm
 | 
|  |   1561 |                         (false,true,false) (prover used') ctxt' (tych Q)
 | 
|  |   1562 |                       handle Utils.ERR _ => Thm.reflexive (tych Q)
 | 
|  |   1563 |                      val lhs_eeq_lhs2 = implies_intr_list ants1 Q_eeq_Q1
 | 
|  |   1564 |                      val lhs_eq_lhs2 = lhs_eeq_lhs2 RS meta_eq_to_obj_eq
 | 
|  |   1565 |                      val ant_th = forall_intr_list(map tych vlist)lhs_eq_lhs2
 | 
|  |   1566 |                  in
 | 
|  |   1567 |                  ant_th COMP thm
 | 
|  |   1568 |               end end
 | 
|  |   1569 | 
 | 
|  |   1570 |              fun eliminate thm =
 | 
|  |   1571 |                case Thm.prop_of thm of
 | 
|  |   1572 |                  Const(@{const_name Pure.imp},_) $ imp $ _ =>
 | 
|  |   1573 |                    eliminate
 | 
|  |   1574 |                     (if not(is_all imp)
 | 
|  |   1575 |                      then uq_eliminate (thm, imp)
 | 
|  |   1576 |                      else q_eliminate (thm, imp))
 | 
|  |   1577 |                             (* Assume that the leading constant is ==,   *)
 | 
|  |   1578 |                 | _ => thm  (* if it is not a ==>                        *)
 | 
|  |   1579 |          in SOME(eliminate (rename thm)) end
 | 
|  |   1580 |          handle Utils.ERR _ => NONE    (* FIXME handle THM as well?? *)
 | 
|  |   1581 | 
 | 
|  |   1582 |         fun restrict_prover ctxt thm =
 | 
|  |   1583 |           let val _ = say "restrict_prover:"
 | 
|  |   1584 |               val cntxt = rev (Simplifier.prems_of ctxt)
 | 
|  |   1585 |               val _ = print_thms ctxt "cntxt:" cntxt
 | 
|  |   1586 |               val Const(@{const_name Pure.imp},_) $ (Const(@{const_name Trueprop},_) $ A) $ _ =
 | 
|  |   1587 |                 Thm.prop_of thm
 | 
|  |   1588 |               fun genl tm = let val vlist = subtract (op aconv) globals
 | 
|  |   1589 |                                            (Misc_Legacy.add_term_frees(tm,[]))
 | 
|  |   1590 |                             in fold_rev Forall vlist tm
 | 
|  |   1591 |                             end
 | 
|  |   1592 |               (*--------------------------------------------------------------
 | 
|  |   1593 |                * This actually isn't quite right, since it will think that
 | 
|  |   1594 |                * not-fully applied occs. of "f" in the context mean that the
 | 
|  |   1595 |                * current call is nested. The real solution is to pass in a
 | 
|  |   1596 |                * term "f v1..vn" which is a pattern that any full application
 | 
|  |   1597 |                * of "f" will match.
 | 
|  |   1598 |                *-------------------------------------------------------------*)
 | 
|  |   1599 |               val func_name = #1(dest_Const func)
 | 
|  |   1600 |               fun is_func (Const (name,_)) = (name = func_name)
 | 
|  |   1601 |                 | is_func _                = false
 | 
|  |   1602 |               val rcontext = rev cntxt
 | 
|  |   1603 |               val cncl = HOLogic.dest_Trueprop o Thm.prop_of
 | 
|  |   1604 |               val antl = case rcontext of [] => []
 | 
|  |   1605 |                          | _   => [USyntax.list_mk_conj(map cncl rcontext)]
 | 
|  |   1606 |               val TC = genl(USyntax.list_mk_imp(antl, A))
 | 
|  |   1607 |               val _ = print_term ctxt "func:" func
 | 
|  |   1608 |               val _ = print_term ctxt "TC:" (HOLogic.mk_Trueprop TC)
 | 
|  |   1609 |               val _ = tc_list := (TC :: !tc_list)
 | 
|  |   1610 |               val nestedp = is_some (USyntax.find_term is_func TC)
 | 
|  |   1611 |               val _ = if nestedp then say "nested" else say "not_nested"
 | 
|  |   1612 |               val th' = if nestedp then raise RULES_ERR "solver" "nested function"
 | 
|  |   1613 |                         else let val cTC = Thm.cterm_of ctxt (HOLogic.mk_Trueprop TC)
 | 
|  |   1614 |                              in case rcontext of
 | 
|  |   1615 |                                 [] => SPEC_ALL(ASSUME cTC)
 | 
|  |   1616 |                                | _ => MP (SPEC_ALL (ASSUME cTC))
 | 
|  |   1617 |                                          (LIST_CONJ rcontext)
 | 
|  |   1618 |                              end
 | 
|  |   1619 |               val th'' = th' RS thm
 | 
|  |   1620 |           in SOME (th'')
 | 
|  |   1621 |           end handle Utils.ERR _ => NONE    (* FIXME handle THM as well?? *)
 | 
|  |   1622 |     in
 | 
|  |   1623 |     (if (is_cong thm) then cong_prover else restrict_prover) ctxt thm
 | 
|  |   1624 |     end
 | 
|  |   1625 |     val ctm = Thm.cprop_of th
 | 
|  |   1626 |     val names = Misc_Legacy.add_term_names (Thm.term_of ctm, [])
 | 
|  |   1627 |     val th1 =
 | 
|  |   1628 |       Raw_Simplifier.rewrite_cterm (false, true, false)
 | 
|  |   1629 |         (prover names) (ctxt0 addsimps [cut_lemma'] |> fold Simplifier.add_eqcong congs) ctm
 | 
|  |   1630 |     val th2 = Thm.equal_elim th1 th
 | 
|  |   1631 |  in
 | 
|  |   1632 |  (th2, filter_out restricted (!tc_list))
 | 
|  |   1633 |  end;
 | 
|  |   1634 | 
 | 
|  |   1635 | 
 | 
|  |   1636 | fun prove ctxt strict (t, tac) =
 | 
|  |   1637 |   let
 | 
|  |   1638 |     val ctxt' = Variable.auto_fixes t ctxt;
 | 
|  |   1639 |   in
 | 
|  |   1640 |     if strict
 | 
|  |   1641 |     then Goal.prove ctxt' [] [] t (K tac)
 | 
|  |   1642 |     else Goal.prove ctxt' [] [] t (K tac)
 | 
|  |   1643 |       handle ERROR msg => (warning msg; raise RULES_ERR "prove" msg)
 | 
|  |   1644 |   end;
 | 
|  |   1645 | 
 | 
|  |   1646 | end;
 | 
|  |   1647 | 
 | 
|  |   1648 | 
 | 
| 60521 |   1649 | 
 | 
| 60520 |   1650 | (*** theory operations ***)
 | 
|  |   1651 | 
 | 
|  |   1652 | structure Thry: THRY =
 | 
|  |   1653 | struct
 | 
|  |   1654 | 
 | 
|  |   1655 | 
 | 
|  |   1656 | fun THRY_ERR func mesg = Utils.ERR {module = "Thry", func = func, mesg = mesg};
 | 
|  |   1657 | 
 | 
|  |   1658 | 
 | 
|  |   1659 | (*---------------------------------------------------------------------------
 | 
|  |   1660 |  *    Matching
 | 
|  |   1661 |  *---------------------------------------------------------------------------*)
 | 
|  |   1662 | 
 | 
|  |   1663 | local
 | 
|  |   1664 | 
 | 
|  |   1665 | fun tybind (ixn, (S, T)) = (TVar (ixn, S), T);
 | 
|  |   1666 | 
 | 
|  |   1667 | in
 | 
|  |   1668 | 
 | 
|  |   1669 | fun match_term thry pat ob =
 | 
|  |   1670 |   let
 | 
|  |   1671 |     val (ty_theta, tm_theta) = Pattern.match thry (pat,ob) (Vartab.empty, Vartab.empty);
 | 
|  |   1672 |     fun tmbind (ixn, (T, t)) = (Var (ixn, Envir.subst_type ty_theta T), t)
 | 
|  |   1673 |   in (map tmbind (Vartab.dest tm_theta), map tybind (Vartab.dest ty_theta))
 | 
|  |   1674 |   end;
 | 
|  |   1675 | 
 | 
|  |   1676 | fun match_type thry pat ob =
 | 
|  |   1677 |   map tybind (Vartab.dest (Sign.typ_match thry (pat, ob) Vartab.empty));
 | 
|  |   1678 | 
 | 
|  |   1679 | end;
 | 
|  |   1680 | 
 | 
|  |   1681 | 
 | 
|  |   1682 | (*---------------------------------------------------------------------------
 | 
|  |   1683 |  * Typing
 | 
|  |   1684 |  *---------------------------------------------------------------------------*)
 | 
|  |   1685 | 
 | 
|  |   1686 | fun typecheck thy t =
 | 
|  |   1687 |   Thm.global_cterm_of thy t
 | 
|  |   1688 |     handle TYPE (msg, _, _) => raise THRY_ERR "typecheck" msg
 | 
|  |   1689 |       | TERM (msg, _) => raise THRY_ERR "typecheck" msg;
 | 
|  |   1690 | 
 | 
|  |   1691 | 
 | 
|  |   1692 | (*---------------------------------------------------------------------------
 | 
|  |   1693 |  * Get information about datatypes
 | 
|  |   1694 |  *---------------------------------------------------------------------------*)
 | 
|  |   1695 | 
 | 
|  |   1696 | fun match_info thy dtco =
 | 
|  |   1697 |   case (BNF_LFP_Compat.get_info thy [BNF_LFP_Compat.Keep_Nesting] dtco,
 | 
|  |   1698 |          BNF_LFP_Compat.get_constrs thy dtco) of
 | 
|  |   1699 |       (SOME {case_name, ... }, SOME constructors) =>
 | 
|  |   1700 |         SOME {case_const = Const (case_name, Sign.the_const_type thy case_name), constructors = map Const constructors}
 | 
|  |   1701 |     | _ => NONE;
 | 
|  |   1702 | 
 | 
|  |   1703 | fun induct_info thy dtco = case BNF_LFP_Compat.get_info thy [BNF_LFP_Compat.Keep_Nesting] dtco of
 | 
|  |   1704 |         NONE => NONE
 | 
|  |   1705 |       | SOME {nchotomy, ...} =>
 | 
|  |   1706 |           SOME {nchotomy = nchotomy,
 | 
|  |   1707 |                 constructors = (map Const o the o BNF_LFP_Compat.get_constrs thy) dtco};
 | 
|  |   1708 | 
 | 
|  |   1709 | fun extract_info thy =
 | 
|  |   1710 |  let val infos = map snd (Symtab.dest (BNF_LFP_Compat.get_all thy [BNF_LFP_Compat.Keep_Nesting]))
 | 
|  |   1711 |  in {case_congs = map (mk_meta_eq o #case_cong) infos,
 | 
|  |   1712 |      case_rewrites = maps (map mk_meta_eq o #case_rewrites) infos}
 | 
|  |   1713 |  end;
 | 
|  |   1714 | 
 | 
|  |   1715 | 
 | 
|  |   1716 | end;
 | 
|  |   1717 | 
 | 
|  |   1718 | 
 | 
| 60521 |   1719 | 
 | 
| 60520 |   1720 | (*** first part of main module ***)
 | 
|  |   1721 | 
 | 
|  |   1722 | structure Prim: PRIM =
 | 
|  |   1723 | struct
 | 
|  |   1724 | 
 | 
|  |   1725 | val trace = Unsynchronized.ref false;
 | 
|  |   1726 | 
 | 
|  |   1727 | 
 | 
|  |   1728 | fun TFL_ERR func mesg = Utils.ERR {module = "Tfl", func = func, mesg = mesg};
 | 
|  |   1729 | 
 | 
|  |   1730 | val concl = #2 o Rules.dest_thm;
 | 
|  |   1731 | 
 | 
|  |   1732 | val list_mk_type = Utils.end_itlist (curry (op -->));
 | 
|  |   1733 | 
 | 
|  |   1734 | fun front_last [] = raise TFL_ERR "front_last" "empty list"
 | 
|  |   1735 |   | front_last [x] = ([],x)
 | 
|  |   1736 |   | front_last (h::t) =
 | 
|  |   1737 |      let val (pref,x) = front_last t
 | 
|  |   1738 |      in
 | 
|  |   1739 |         (h::pref,x)
 | 
|  |   1740 |      end;
 | 
|  |   1741 | 
 | 
|  |   1742 | 
 | 
|  |   1743 | (*---------------------------------------------------------------------------
 | 
|  |   1744 |  * The next function is common to pattern-match translation and
 | 
|  |   1745 |  * proof of completeness of cases for the induction theorem.
 | 
|  |   1746 |  *
 | 
|  |   1747 |  * The curried function "gvvariant" returns a function to generate distinct
 | 
|  |   1748 |  * variables that are guaranteed not to be in names.  The names of
 | 
|  |   1749 |  * the variables go u, v, ..., z, aa, ..., az, ...  The returned
 | 
|  |   1750 |  * function contains embedded refs!
 | 
|  |   1751 |  *---------------------------------------------------------------------------*)
 | 
|  |   1752 | fun gvvariant names =
 | 
|  |   1753 |   let val slist = Unsynchronized.ref names
 | 
|  |   1754 |       val vname = Unsynchronized.ref "u"
 | 
|  |   1755 |       fun new() =
 | 
|  |   1756 |          if member (op =) (!slist) (!vname)
 | 
|  |   1757 |          then (vname := Symbol.bump_string (!vname);  new())
 | 
|  |   1758 |          else (slist := !vname :: !slist;  !vname)
 | 
|  |   1759 |   in
 | 
|  |   1760 |   fn ty => Free(new(), ty)
 | 
|  |   1761 |   end;
 | 
|  |   1762 | 
 | 
|  |   1763 | 
 | 
|  |   1764 | (*---------------------------------------------------------------------------
 | 
|  |   1765 |  * Used in induction theorem production. This is the simple case of
 | 
|  |   1766 |  * partitioning up pattern rows by the leading constructor.
 | 
|  |   1767 |  *---------------------------------------------------------------------------*)
 | 
|  |   1768 | fun ipartition gv (constructors,rows) =
 | 
|  |   1769 |   let fun pfail s = raise TFL_ERR "partition.part" s
 | 
|  |   1770 |       fun part {constrs = [],   rows = [],   A} = rev A
 | 
|  |   1771 |         | part {constrs = [],   rows = _::_, A} = pfail"extra cases in defn"
 | 
|  |   1772 |         | part {constrs = _::_, rows = [],   A} = pfail"cases missing in defn"
 | 
|  |   1773 |         | part {constrs = c::crst, rows,     A} =
 | 
|  |   1774 |           let val (c, T) = dest_Const c
 | 
|  |   1775 |               val L = binder_types T
 | 
|  |   1776 |               val (in_group, not_in_group) =
 | 
|  |   1777 |                fold_rev (fn (row as (p::rst, rhs)) =>
 | 
|  |   1778 |                          fn (in_group,not_in_group) =>
 | 
|  |   1779 |                   let val (pc,args) = USyntax.strip_comb p
 | 
|  |   1780 |                   in if (#1(dest_Const pc) = c)
 | 
|  |   1781 |                      then ((args@rst, rhs)::in_group, not_in_group)
 | 
|  |   1782 |                      else (in_group, row::not_in_group)
 | 
|  |   1783 |                   end)      rows ([],[])
 | 
|  |   1784 |               val col_types = Utils.take type_of (length L, #1(hd in_group))
 | 
|  |   1785 |           in
 | 
|  |   1786 |           part{constrs = crst, rows = not_in_group,
 | 
|  |   1787 |                A = {constructor = c,
 | 
|  |   1788 |                     new_formals = map gv col_types,
 | 
|  |   1789 |                     group = in_group}::A}
 | 
|  |   1790 |           end
 | 
|  |   1791 |   in part{constrs = constructors, rows = rows, A = []}
 | 
|  |   1792 |   end;
 | 
|  |   1793 | 
 | 
|  |   1794 | 
 | 
|  |   1795 | 
 | 
|  |   1796 | (*---------------------------------------------------------------------------
 | 
|  |   1797 |  * Each pattern carries with it a tag (i,b) where
 | 
|  |   1798 |  * i is the clause it came from and
 | 
|  |   1799 |  * b=true indicates that clause was given by the user
 | 
|  |   1800 |  * (or is an instantiation of a user supplied pattern)
 | 
|  |   1801 |  * b=false --> i = ~1
 | 
|  |   1802 |  *---------------------------------------------------------------------------*)
 | 
|  |   1803 | 
 | 
|  |   1804 | type pattern = term * (int * bool)
 | 
|  |   1805 | 
 | 
|  |   1806 | fun pattern_map f (tm,x) = (f tm, x);
 | 
|  |   1807 | 
 | 
|  |   1808 | fun pattern_subst theta = pattern_map (subst_free theta);
 | 
|  |   1809 | 
 | 
|  |   1810 | val pat_of = fst;
 | 
|  |   1811 | fun row_of_pat x = fst (snd x);
 | 
|  |   1812 | fun given x = snd (snd x);
 | 
|  |   1813 | 
 | 
|  |   1814 | (*---------------------------------------------------------------------------
 | 
|  |   1815 |  * Produce an instance of a constructor, plus genvars for its arguments.
 | 
|  |   1816 |  *---------------------------------------------------------------------------*)
 | 
|  |   1817 | fun fresh_constr ty_match colty gv c =
 | 
|  |   1818 |   let val (_,Ty) = dest_Const c
 | 
|  |   1819 |       val L = binder_types Ty
 | 
|  |   1820 |       and ty = body_type Ty
 | 
|  |   1821 |       val ty_theta = ty_match ty colty
 | 
|  |   1822 |       val c' = USyntax.inst ty_theta c
 | 
|  |   1823 |       val gvars = map (USyntax.inst ty_theta o gv) L
 | 
|  |   1824 |   in (c', gvars)
 | 
|  |   1825 |   end;
 | 
|  |   1826 | 
 | 
|  |   1827 | 
 | 
|  |   1828 | (*---------------------------------------------------------------------------
 | 
|  |   1829 |  * Goes through a list of rows and picks out the ones beginning with a
 | 
|  |   1830 |  * pattern with constructor = name.
 | 
|  |   1831 |  *---------------------------------------------------------------------------*)
 | 
|  |   1832 | fun mk_group name rows =
 | 
|  |   1833 |   fold_rev (fn (row as ((prfx, p::rst), rhs)) =>
 | 
|  |   1834 |             fn (in_group,not_in_group) =>
 | 
|  |   1835 |                let val (pc,args) = USyntax.strip_comb p
 | 
|  |   1836 |                in if ((#1 (Term.dest_Const pc) = name) handle TERM _ => false)
 | 
|  |   1837 |                   then (((prfx,args@rst), rhs)::in_group, not_in_group)
 | 
|  |   1838 |                   else (in_group, row::not_in_group) end)
 | 
|  |   1839 |       rows ([],[]);
 | 
|  |   1840 | 
 | 
|  |   1841 | (*---------------------------------------------------------------------------
 | 
|  |   1842 |  * Partition the rows. Not efficient: we should use hashing.
 | 
|  |   1843 |  *---------------------------------------------------------------------------*)
 | 
|  |   1844 | fun partition _ _ (_,_,_,[]) = raise TFL_ERR "partition" "no rows"
 | 
|  |   1845 |   | partition gv ty_match
 | 
|  |   1846 |               (constructors, colty, res_ty, rows as (((prfx,_),_)::_)) =
 | 
|  |   1847 | let val fresh = fresh_constr ty_match colty gv
 | 
|  |   1848 |      fun part {constrs = [],      rows, A} = rev A
 | 
|  |   1849 |        | part {constrs = c::crst, rows, A} =
 | 
|  |   1850 |          let val (c',gvars) = fresh c
 | 
|  |   1851 |              val (in_group, not_in_group) = mk_group (#1 (dest_Const c')) rows
 | 
|  |   1852 |              val in_group' =
 | 
|  |   1853 |                  if (null in_group)  (* Constructor not given *)
 | 
|  |   1854 |                  then [((prfx, #2(fresh c)), (USyntax.ARB res_ty, (~1,false)))]
 | 
|  |   1855 |                  else in_group
 | 
|  |   1856 |          in
 | 
|  |   1857 |          part{constrs = crst,
 | 
|  |   1858 |               rows = not_in_group,
 | 
|  |   1859 |               A = {constructor = c',
 | 
|  |   1860 |                    new_formals = gvars,
 | 
|  |   1861 |                    group = in_group'}::A}
 | 
|  |   1862 |          end
 | 
|  |   1863 | in part{constrs=constructors, rows=rows, A=[]}
 | 
|  |   1864 | end;
 | 
|  |   1865 | 
 | 
|  |   1866 | (*---------------------------------------------------------------------------
 | 
|  |   1867 |  * Misc. routines used in mk_case
 | 
|  |   1868 |  *---------------------------------------------------------------------------*)
 | 
|  |   1869 | 
 | 
|  |   1870 | fun mk_pat (c,l) =
 | 
|  |   1871 |   let val L = length (binder_types (type_of c))
 | 
|  |   1872 |       fun build (prfx,tag,plist) =
 | 
|  |   1873 |           let val (args, plist') = chop L plist
 | 
|  |   1874 |           in (prfx,tag,list_comb(c,args)::plist') end
 | 
|  |   1875 |   in map build l end;
 | 
|  |   1876 | 
 | 
|  |   1877 | fun v_to_prfx (prfx, v::pats) = (v::prfx,pats)
 | 
|  |   1878 |   | v_to_prfx _ = raise TFL_ERR "mk_case" "v_to_prfx";
 | 
|  |   1879 | 
 | 
|  |   1880 | fun v_to_pats (v::prfx,tag, pats) = (prfx, tag, v::pats)
 | 
|  |   1881 |   | v_to_pats _ = raise TFL_ERR "mk_case" "v_to_pats";
 | 
|  |   1882 | 
 | 
|  |   1883 | 
 | 
|  |   1884 | (*----------------------------------------------------------------------------
 | 
|  |   1885 |  * Translation of pattern terms into nested case expressions.
 | 
|  |   1886 |  *
 | 
|  |   1887 |  * This performs the translation and also builds the full set of patterns.
 | 
|  |   1888 |  * Thus it supports the construction of induction theorems even when an
 | 
|  |   1889 |  * incomplete set of patterns is given.
 | 
|  |   1890 |  *---------------------------------------------------------------------------*)
 | 
|  |   1891 | 
 | 
|  |   1892 | fun mk_case ty_info ty_match usednames range_ty =
 | 
|  |   1893 |  let
 | 
|  |   1894 |  fun mk_case_fail s = raise TFL_ERR "mk_case" s
 | 
|  |   1895 |  val fresh_var = gvvariant usednames
 | 
|  |   1896 |  val divide = partition fresh_var ty_match
 | 
| 60521 |   1897 |  fun expand _ ty ((_,[]), _) = mk_case_fail"expand_var_row"
 | 
| 60520 |   1898 |    | expand constructors ty (row as ((prfx, p::rst), rhs)) =
 | 
|  |   1899 |        if (is_Free p)
 | 
|  |   1900 |        then let val fresh = fresh_constr ty_match ty fresh_var
 | 
|  |   1901 |                 fun expnd (c,gvs) =
 | 
|  |   1902 |                   let val capp = list_comb(c,gvs)
 | 
|  |   1903 |                   in ((prfx, capp::rst), pattern_subst[(p,capp)] rhs)
 | 
|  |   1904 |                   end
 | 
|  |   1905 |             in map expnd (map fresh constructors)  end
 | 
|  |   1906 |        else [row]
 | 
|  |   1907 |  fun mk{rows=[],...} = mk_case_fail"no rows"
 | 
|  |   1908 |    | mk{path=[], rows = ((prfx, []), (tm,tag))::_} =  (* Done *)
 | 
|  |   1909 |         ([(prfx,tag,[])], tm)
 | 
|  |   1910 |    | mk{path=[], rows = _::_} = mk_case_fail"blunder"
 | 
|  |   1911 |    | mk{path as u::rstp, rows as ((prfx, []), rhs)::rst} =
 | 
|  |   1912 |         mk{path = path,
 | 
|  |   1913 |            rows = ((prfx, [fresh_var(type_of u)]), rhs)::rst}
 | 
|  |   1914 |    | mk{path = u::rstp, rows as ((_, p::_), _)::_} =
 | 
|  |   1915 |      let val (pat_rectangle,rights) = ListPair.unzip rows
 | 
|  |   1916 |          val col0 = map(hd o #2) pat_rectangle
 | 
|  |   1917 |      in
 | 
|  |   1918 |      if (forall is_Free col0)
 | 
|  |   1919 |      then let val rights' = map (fn(v,e) => pattern_subst[(v,u)] e)
 | 
|  |   1920 |                                 (ListPair.zip (col0, rights))
 | 
|  |   1921 |               val pat_rectangle' = map v_to_prfx pat_rectangle
 | 
|  |   1922 |               val (pref_patl,tm) = mk{path = rstp,
 | 
|  |   1923 |                                       rows = ListPair.zip (pat_rectangle',
 | 
|  |   1924 |                                                            rights')}
 | 
|  |   1925 |           in (map v_to_pats pref_patl, tm)
 | 
|  |   1926 |           end
 | 
|  |   1927 |      else
 | 
|  |   1928 |      let val pty as Type (ty_name,_) = type_of p
 | 
|  |   1929 |      in
 | 
|  |   1930 |      case (ty_info ty_name)
 | 
|  |   1931 |      of NONE => mk_case_fail("Not a known datatype: "^ty_name)
 | 
|  |   1932 |       | SOME{case_const,constructors} =>
 | 
|  |   1933 |         let
 | 
|  |   1934 |             val case_const_name = #1(dest_Const case_const)
 | 
|  |   1935 |             val nrows = maps (expand constructors pty) rows
 | 
|  |   1936 |             val subproblems = divide(constructors, pty, range_ty, nrows)
 | 
|  |   1937 |             val groups      = map #group subproblems
 | 
|  |   1938 |             and new_formals = map #new_formals subproblems
 | 
|  |   1939 |             and constructors' = map #constructor subproblems
 | 
|  |   1940 |             val news = map (fn (nf,rows) => {path = nf@rstp, rows=rows})
 | 
|  |   1941 |                            (ListPair.zip (new_formals, groups))
 | 
|  |   1942 |             val rec_calls = map mk news
 | 
|  |   1943 |             val (pat_rect,dtrees) = ListPair.unzip rec_calls
 | 
|  |   1944 |             val case_functions = map USyntax.list_mk_abs
 | 
|  |   1945 |                                   (ListPair.zip (new_formals, dtrees))
 | 
|  |   1946 |             val types = map type_of (case_functions@[u]) @ [range_ty]
 | 
|  |   1947 |             val case_const' = Const(case_const_name, list_mk_type types)
 | 
|  |   1948 |             val tree = list_comb(case_const', case_functions@[u])
 | 
|  |   1949 |             val pat_rect1 = flat (ListPair.map mk_pat (constructors', pat_rect))
 | 
|  |   1950 |         in (pat_rect1,tree)
 | 
|  |   1951 |         end
 | 
|  |   1952 |      end end
 | 
|  |   1953 |  in mk
 | 
|  |   1954 |  end;
 | 
|  |   1955 | 
 | 
|  |   1956 | 
 | 
|  |   1957 | (* Repeated variable occurrences in a pattern are not allowed. *)
 | 
|  |   1958 | fun FV_multiset tm =
 | 
|  |   1959 |    case (USyntax.dest_term tm)
 | 
|  |   1960 |      of USyntax.VAR{Name = c, Ty = T} => [Free(c, T)]
 | 
|  |   1961 |       | USyntax.CONST _ => []
 | 
|  |   1962 |       | USyntax.COMB{Rator, Rand} => FV_multiset Rator @ FV_multiset Rand
 | 
|  |   1963 |       | USyntax.LAMB _ => raise TFL_ERR "FV_multiset" "lambda";
 | 
|  |   1964 | 
 | 
|  |   1965 | fun no_repeat_vars thy pat =
 | 
|  |   1966 |  let fun check [] = true
 | 
|  |   1967 |        | check (v::rst) =
 | 
|  |   1968 |          if member (op aconv) rst v then
 | 
|  |   1969 |             raise TFL_ERR "no_repeat_vars"
 | 
|  |   1970 |                           (quote (#1 (dest_Free v)) ^
 | 
|  |   1971 |                           " occurs repeatedly in the pattern " ^
 | 
|  |   1972 |                           quote (Syntax.string_of_term_global thy pat))
 | 
|  |   1973 |          else check rst
 | 
|  |   1974 |  in check (FV_multiset pat)
 | 
|  |   1975 |  end;
 | 
|  |   1976 | 
 | 
|  |   1977 | fun dest_atom (Free p) = p
 | 
|  |   1978 |   | dest_atom (Const p) = p
 | 
|  |   1979 |   | dest_atom  _ = raise TFL_ERR "dest_atom" "function name not an identifier";
 | 
|  |   1980 | 
 | 
|  |   1981 | fun same_name (p,q) = #1(dest_atom p) = #1(dest_atom q);
 | 
|  |   1982 | 
 | 
|  |   1983 | local fun mk_functional_err s = raise TFL_ERR "mk_functional" s
 | 
|  |   1984 |       fun single [_$_] =
 | 
|  |   1985 |               mk_functional_err "recdef does not allow currying"
 | 
|  |   1986 |         | single [f] = f
 | 
|  |   1987 |         | single fs  =
 | 
|  |   1988 |               (*multiple function names?*)
 | 
|  |   1989 |               if length (distinct same_name fs) < length fs
 | 
|  |   1990 |               then mk_functional_err
 | 
|  |   1991 |                    "The function being declared appears with multiple types"
 | 
|  |   1992 |               else mk_functional_err
 | 
|  |   1993 |                    (string_of_int (length fs) ^
 | 
|  |   1994 |                     " distinct function names being declared")
 | 
|  |   1995 | in
 | 
|  |   1996 | fun mk_functional thy clauses =
 | 
|  |   1997 |  let val (L,R) = ListPair.unzip (map HOLogic.dest_eq clauses
 | 
|  |   1998 |                    handle TERM _ => raise TFL_ERR "mk_functional"
 | 
|  |   1999 |                         "recursion equations must use the = relation")
 | 
|  |   2000 |      val (funcs,pats) = ListPair.unzip (map (fn (t$u) =>(t,u)) L)
 | 
|  |   2001 |      val atom = single (distinct (op aconv) funcs)
 | 
|  |   2002 |      val (fname,ftype) = dest_atom atom
 | 
| 60521 |   2003 |      val _ = map (no_repeat_vars thy) pats
 | 
| 60520 |   2004 |      val rows = ListPair.zip (map (fn x => ([]:term list,[x])) pats,
 | 
|  |   2005 |                               map_index (fn (i, t) => (t,(i,true))) R)
 | 
|  |   2006 |      val names = List.foldr Misc_Legacy.add_term_names [] R
 | 
|  |   2007 |      val atype = type_of(hd pats)
 | 
|  |   2008 |      and aname = singleton (Name.variant_list names) "a"
 | 
|  |   2009 |      val a = Free(aname,atype)
 | 
|  |   2010 |      val ty_info = Thry.match_info thy
 | 
|  |   2011 |      val ty_match = Thry.match_type thy
 | 
|  |   2012 |      val range_ty = type_of (hd R)
 | 
|  |   2013 |      val (patts, case_tm) = mk_case ty_info ty_match (aname::names) range_ty
 | 
|  |   2014 |                                     {path=[a], rows=rows}
 | 
|  |   2015 |      val patts1 = map (fn (_,tag,[pat]) => (pat,tag)) patts
 | 
|  |   2016 |           handle Match => mk_functional_err "error in pattern-match translation"
 | 
|  |   2017 |      val patts2 = Library.sort (Library.int_ord o apply2 row_of_pat) patts1
 | 
|  |   2018 |      val finals = map row_of_pat patts2
 | 
|  |   2019 |      val originals = map (row_of_pat o #2) rows
 | 
| 60521 |   2020 |      val _ = case (subtract (op =) finals originals)
 | 
| 60520 |   2021 |              of [] => ()
 | 
|  |   2022 |           | L => mk_functional_err
 | 
|  |   2023 |  ("The following clauses are redundant (covered by preceding clauses): " ^
 | 
|  |   2024 |                    commas (map (fn i => string_of_int (i + 1)) L))
 | 
|  |   2025 |  in {functional = Abs(Long_Name.base_name fname, ftype,
 | 
|  |   2026 |                       abstract_over (atom, absfree (aname,atype) case_tm)),
 | 
|  |   2027 |      pats = patts2}
 | 
|  |   2028 | end end;
 | 
|  |   2029 | 
 | 
|  |   2030 | 
 | 
|  |   2031 | (*----------------------------------------------------------------------------
 | 
|  |   2032 |  *
 | 
|  |   2033 |  *                    PRINCIPLES OF DEFINITION
 | 
|  |   2034 |  *
 | 
|  |   2035 |  *---------------------------------------------------------------------------*)
 | 
|  |   2036 | 
 | 
|  |   2037 | 
 | 
|  |   2038 | (*For Isabelle, the lhs of a definition must be a constant.*)
 | 
|  |   2039 | fun const_def sign (c, Ty, rhs) =
 | 
|  |   2040 |   singleton (Syntax.check_terms (Proof_Context.init_global sign))
 | 
|  |   2041 |     (Const(@{const_name Pure.eq},dummyT) $ Const(c,Ty) $ rhs);
 | 
|  |   2042 | 
 | 
|  |   2043 | (*Make all TVars available for instantiation by adding a ? to the front*)
 | 
|  |   2044 | fun poly_tvars (Type(a,Ts)) = Type(a, map (poly_tvars) Ts)
 | 
|  |   2045 |   | poly_tvars (TFree (a,sort)) = TVar (("?" ^ a, 0), sort)
 | 
|  |   2046 |   | poly_tvars (TVar ((a,i),sort)) = TVar (("?" ^ a, i+1), sort);
 | 
|  |   2047 | 
 | 
|  |   2048 | local
 | 
|  |   2049 |   val f_eq_wfrec_R_M =
 | 
| 60522 |   2050 |     #ant(USyntax.dest_imp(#2(USyntax.strip_forall (concl @{thm tfl_wfrec}))))
 | 
| 60520 |   2051 |   val {lhs=f, rhs} = USyntax.dest_eq f_eq_wfrec_R_M
 | 
| 60521 |   2052 |   val _ = dest_Free f
 | 
| 60520 |   2053 |   val (wfrec,_) = USyntax.strip_comb rhs
 | 
|  |   2054 | in
 | 
|  |   2055 | 
 | 
|  |   2056 | fun wfrec_definition0 fid R (functional as Abs(x, Ty, _)) thy =
 | 
|  |   2057 |   let
 | 
|  |   2058 |     val def_name = Thm.def_name (Long_Name.base_name fid)
 | 
|  |   2059 |     val wfrec_R_M = map_types poly_tvars (wfrec $ map_types poly_tvars R) $ functional
 | 
|  |   2060 |     val def_term = const_def thy (fid, Ty, wfrec_R_M)
 | 
|  |   2061 |     val ([def], thy') =
 | 
|  |   2062 |       Global_Theory.add_defs false [Thm.no_attributes (Binding.name def_name, def_term)] thy
 | 
|  |   2063 |   in (def, thy') end;
 | 
|  |   2064 | 
 | 
|  |   2065 | end;
 | 
|  |   2066 | 
 | 
|  |   2067 | 
 | 
|  |   2068 | 
 | 
|  |   2069 | (*---------------------------------------------------------------------------
 | 
|  |   2070 |  * This structure keeps track of congruence rules that aren't derived
 | 
|  |   2071 |  * from a datatype definition.
 | 
|  |   2072 |  *---------------------------------------------------------------------------*)
 | 
|  |   2073 | fun extraction_thms thy =
 | 
|  |   2074 |  let val {case_rewrites,case_congs} = Thry.extract_info thy
 | 
|  |   2075 |  in (case_rewrites, case_congs)
 | 
|  |   2076 |  end;
 | 
|  |   2077 | 
 | 
|  |   2078 | 
 | 
|  |   2079 | (*---------------------------------------------------------------------------
 | 
|  |   2080 |  * Pair patterns with termination conditions. The full list of patterns for
 | 
|  |   2081 |  * a definition is merged with the TCs arising from the user-given clauses.
 | 
|  |   2082 |  * There can be fewer clauses than the full list, if the user omitted some
 | 
|  |   2083 |  * cases. This routine is used to prepare input for mk_induction.
 | 
|  |   2084 |  *---------------------------------------------------------------------------*)
 | 
|  |   2085 | fun merge full_pats TCs =
 | 
|  |   2086 | let fun insert (p,TCs) =
 | 
|  |   2087 |       let fun insrt ((x as (h,[]))::rst) =
 | 
|  |   2088 |                  if (p aconv h) then (p,TCs)::rst else x::insrt rst
 | 
|  |   2089 |             | insrt (x::rst) = x::insrt rst
 | 
|  |   2090 |             | insrt[] = raise TFL_ERR "merge.insert" "pattern not found"
 | 
|  |   2091 |       in insrt end
 | 
|  |   2092 |     fun pass ([],ptcl_final) = ptcl_final
 | 
|  |   2093 |       | pass (ptcs::tcl, ptcl) = pass(tcl, insert ptcs ptcl)
 | 
|  |   2094 | in
 | 
|  |   2095 |   pass (TCs, map (fn p => (p,[])) full_pats)
 | 
|  |   2096 | end;
 | 
|  |   2097 | 
 | 
|  |   2098 | 
 | 
|  |   2099 | fun givens pats = map pat_of (filter given pats);
 | 
|  |   2100 | 
 | 
|  |   2101 | fun post_definition ctxt meta_tflCongs (def, pats) =
 | 
|  |   2102 |  let val thy = Proof_Context.theory_of ctxt
 | 
|  |   2103 |      val tych = Thry.typecheck thy
 | 
|  |   2104 |      val f = #lhs(USyntax.dest_eq(concl def))
 | 
| 60522 |   2105 |      val corollary = Rules.MATCH_MP @{thm tfl_wfrec} def
 | 
| 60520 |   2106 |      val pats' = filter given pats
 | 
|  |   2107 |      val given_pats = map pat_of pats'
 | 
|  |   2108 |      val rows = map row_of_pat pats'
 | 
|  |   2109 |      val WFR = #ant(USyntax.dest_imp(concl corollary))
 | 
|  |   2110 |      val R = #Rand(USyntax.dest_comb WFR)
 | 
|  |   2111 |      val corollary' = Rules.UNDISCH corollary  (* put WF R on assums *)
 | 
|  |   2112 |      val corollaries = map (fn pat => Rules.SPEC (tych pat) corollary') given_pats
 | 
|  |   2113 |      val (case_rewrites,context_congs) = extraction_thms thy
 | 
|  |   2114 |      (*case_ss causes minimal simplification: bodies of case expressions are
 | 
|  |   2115 |        not simplified. Otherwise large examples (Red-Black trees) are too
 | 
|  |   2116 |        slow.*)
 | 
|  |   2117 |      val case_simpset =
 | 
|  |   2118 |        put_simpset HOL_basic_ss ctxt
 | 
|  |   2119 |           addsimps case_rewrites
 | 
|  |   2120 |           |> fold (Simplifier.add_cong o #case_cong_weak o snd)
 | 
|  |   2121 |               (Symtab.dest (BNF_LFP_Compat.get_all thy [BNF_LFP_Compat.Keep_Nesting]))
 | 
|  |   2122 |      val corollaries' = map (Simplifier.simplify case_simpset) corollaries
 | 
|  |   2123 |      val extract =
 | 
|  |   2124 |       Rules.CONTEXT_REWRITE_RULE ctxt (f, [R], @{thm cut_apply}, meta_tflCongs @ context_congs)
 | 
|  |   2125 |      val (rules, TCs) = ListPair.unzip (map extract corollaries')
 | 
| 60522 |   2126 |      val rules0 = map (rewrite_rule ctxt @{thms tfl_cut_def}) rules
 | 
| 60520 |   2127 |      val mk_cond_rule = Rules.FILTER_DISCH_ALL(not o curry (op aconv) WFR)
 | 
|  |   2128 |      val rules1 = Rules.LIST_CONJ(map mk_cond_rule rules0)
 | 
|  |   2129 |  in
 | 
|  |   2130 |  {rules = rules1,
 | 
|  |   2131 |   rows = rows,
 | 
|  |   2132 |   full_pats_TCs = merge (map pat_of pats) (ListPair.zip (given_pats, TCs)),
 | 
|  |   2133 |   TCs = TCs}
 | 
|  |   2134 |  end;
 | 
|  |   2135 | 
 | 
|  |   2136 | 
 | 
|  |   2137 | (*---------------------------------------------------------------------------
 | 
|  |   2138 |  * Perform the extraction without making the definition. Definition and
 | 
|  |   2139 |  * extraction commute for the non-nested case.  (Deferred recdefs)
 | 
|  |   2140 |  *
 | 
|  |   2141 |  * The purpose of wfrec_eqns is merely to instantiate the recursion theorem
 | 
|  |   2142 |  * and extract termination conditions: no definition is made.
 | 
|  |   2143 |  *---------------------------------------------------------------------------*)
 | 
|  |   2144 | 
 | 
|  |   2145 | fun wfrec_eqns thy fid tflCongs eqns =
 | 
|  |   2146 |  let val ctxt = Proof_Context.init_global thy
 | 
|  |   2147 |      val {lhs,rhs} = USyntax.dest_eq (hd eqns)
 | 
|  |   2148 |      val (f,args) = USyntax.strip_comb lhs
 | 
| 60521 |   2149 |      val (fname,_) = dest_atom f
 | 
|  |   2150 |      val (SV,_) = front_last args    (* SV = schematic variables *)
 | 
| 60520 |   2151 |      val g = list_comb(f,SV)
 | 
|  |   2152 |      val h = Free(fname,type_of g)
 | 
|  |   2153 |      val eqns1 = map (subst_free[(g,h)]) eqns
 | 
|  |   2154 |      val {functional as Abs(x, Ty, _),  pats} = mk_functional thy eqns1
 | 
|  |   2155 |      val given_pats = givens pats
 | 
|  |   2156 |      (* val f = Free(x,Ty) *)
 | 
|  |   2157 |      val Type("fun", [f_dty, f_rty]) = Ty
 | 
| 60521 |   2158 |      val _ = if x<>fid then
 | 
| 60520 |   2159 |                         raise TFL_ERR "wfrec_eqns"
 | 
|  |   2160 |                                       ("Expected a definition of " ^
 | 
|  |   2161 |                                       quote fid ^ " but found one of " ^
 | 
|  |   2162 |                                       quote x)
 | 
|  |   2163 |                  else ()
 | 
|  |   2164 |      val (case_rewrites,context_congs) = extraction_thms thy
 | 
|  |   2165 |      val tych = Thry.typecheck thy
 | 
| 60522 |   2166 |      val WFREC_THM0 = Rules.ISPEC (tych functional) @{thm tfl_wfrec}
 | 
| 60520 |   2167 |      val Const(@{const_name All},_) $ Abs(Rname,Rtype,_) = concl WFREC_THM0
 | 
|  |   2168 |      val R = Free (singleton (Name.variant_list (List.foldr Misc_Legacy.add_term_names [] eqns)) Rname,
 | 
|  |   2169 |                    Rtype)
 | 
|  |   2170 |      val WFREC_THM = Rules.ISPECL [tych R, tych g] WFREC_THM0
 | 
|  |   2171 |      val ([proto_def, WFR],_) = USyntax.strip_imp(concl WFREC_THM)
 | 
| 60521 |   2172 |      val _ =
 | 
| 60520 |   2173 |            if !trace then
 | 
|  |   2174 |                writeln ("ORIGINAL PROTO_DEF: " ^
 | 
|  |   2175 |                           Syntax.string_of_term_global thy proto_def)
 | 
|  |   2176 |            else ()
 | 
|  |   2177 |      val R1 = USyntax.rand WFR
 | 
|  |   2178 |      val corollary' = Rules.UNDISCH (Rules.UNDISCH WFREC_THM)
 | 
|  |   2179 |      val corollaries = map (fn pat => Rules.SPEC (tych pat) corollary') given_pats
 | 
|  |   2180 |      val corollaries' = map (rewrite_rule ctxt case_rewrites) corollaries
 | 
|  |   2181 |      val extract =
 | 
|  |   2182 |       Rules.CONTEXT_REWRITE_RULE ctxt (f, R1::SV, @{thm cut_apply}, tflCongs @ context_congs)
 | 
|  |   2183 |  in {proto_def = proto_def,
 | 
|  |   2184 |      SV=SV,
 | 
|  |   2185 |      WFR=WFR,
 | 
|  |   2186 |      pats=pats,
 | 
|  |   2187 |      extracta = map extract corollaries'}
 | 
|  |   2188 |  end;
 | 
|  |   2189 | 
 | 
|  |   2190 | 
 | 
|  |   2191 | (*----------------------------------------------------------------------------
 | 
|  |   2192 |  *
 | 
|  |   2193 |  *                           INDUCTION THEOREM
 | 
|  |   2194 |  *
 | 
|  |   2195 |  *---------------------------------------------------------------------------*)
 | 
|  |   2196 | 
 | 
|  |   2197 | 
 | 
|  |   2198 | (*------------------------  Miscellaneous function  --------------------------
 | 
|  |   2199 |  *
 | 
|  |   2200 |  *           [x_1,...,x_n]     ?v_1...v_n. M[v_1,...,v_n]
 | 
|  |   2201 |  *     -----------------------------------------------------------
 | 
|  |   2202 |  *     ( M[x_1,...,x_n], [(x_i,?v_1...v_n. M[v_1,...,v_n]),
 | 
|  |   2203 |  *                        ...
 | 
|  |   2204 |  *                        (x_j,?v_n. M[x_1,...,x_(n-1),v_n])] )
 | 
|  |   2205 |  *
 | 
|  |   2206 |  * This function is totally ad hoc. Used in the production of the induction
 | 
|  |   2207 |  * theorem. The nchotomy theorem can have clauses that look like
 | 
|  |   2208 |  *
 | 
|  |   2209 |  *     ?v1..vn. z = C vn..v1
 | 
|  |   2210 |  *
 | 
|  |   2211 |  * in which the order of quantification is not the order of occurrence of the
 | 
|  |   2212 |  * quantified variables as arguments to C. Since we have no control over this
 | 
|  |   2213 |  * aspect of the nchotomy theorem, we make the correspondence explicit by
 | 
|  |   2214 |  * pairing the incoming new variable with the term it gets beta-reduced into.
 | 
|  |   2215 |  *---------------------------------------------------------------------------*)
 | 
|  |   2216 | 
 | 
|  |   2217 | fun alpha_ex_unroll (xlist, tm) =
 | 
|  |   2218 |   let val (qvars,body) = USyntax.strip_exists tm
 | 
|  |   2219 |       val vlist = #2 (USyntax.strip_comb (USyntax.rhs body))
 | 
|  |   2220 |       val plist = ListPair.zip (vlist, xlist)
 | 
|  |   2221 |       val args = map (the o AList.lookup (op aconv) plist) qvars
 | 
|  |   2222 |                    handle Option.Option => raise Fail "TFL.alpha_ex_unroll: no correspondence"
 | 
|  |   2223 |       fun build ex      []   = []
 | 
|  |   2224 |         | build (_$rex) (v::rst) =
 | 
|  |   2225 |            let val ex1 = Term.betapply(rex, v)
 | 
|  |   2226 |            in  ex1 :: build ex1 rst
 | 
|  |   2227 |            end
 | 
|  |   2228 |      val (nex::exl) = rev (tm::build tm args)
 | 
|  |   2229 |   in
 | 
|  |   2230 |   (nex, ListPair.zip (args, rev exl))
 | 
|  |   2231 |   end;
 | 
|  |   2232 | 
 | 
|  |   2233 | 
 | 
|  |   2234 | 
 | 
|  |   2235 | (*----------------------------------------------------------------------------
 | 
|  |   2236 |  *
 | 
|  |   2237 |  *             PROVING COMPLETENESS OF PATTERNS
 | 
|  |   2238 |  *
 | 
|  |   2239 |  *---------------------------------------------------------------------------*)
 | 
|  |   2240 | 
 | 
|  |   2241 | fun mk_case ty_info usednames thy =
 | 
|  |   2242 |  let
 | 
|  |   2243 |  val ctxt = Proof_Context.init_global thy
 | 
|  |   2244 |  val divide = ipartition (gvvariant usednames)
 | 
|  |   2245 |  val tych = Thry.typecheck thy
 | 
|  |   2246 |  fun tych_binding(x,y) = (tych x, tych y)
 | 
|  |   2247 |  fun fail s = raise TFL_ERR "mk_case" s
 | 
|  |   2248 |  fun mk{rows=[],...} = fail"no rows"
 | 
|  |   2249 |    | mk{path=[], rows = [([], (thm, bindings))]} =
 | 
|  |   2250 |                          Rules.IT_EXISTS ctxt (map tych_binding bindings) thm
 | 
|  |   2251 |    | mk{path = u::rstp, rows as (p::_, _)::_} =
 | 
|  |   2252 |      let val (pat_rectangle,rights) = ListPair.unzip rows
 | 
|  |   2253 |          val col0 = map hd pat_rectangle
 | 
|  |   2254 |          val pat_rectangle' = map tl pat_rectangle
 | 
|  |   2255 |      in
 | 
|  |   2256 |      if (forall is_Free col0) (* column 0 is all variables *)
 | 
|  |   2257 |      then let val rights' = map (fn ((thm,theta),v) => (thm,theta@[(u,v)]))
 | 
|  |   2258 |                                 (ListPair.zip (rights, col0))
 | 
|  |   2259 |           in mk{path = rstp, rows = ListPair.zip (pat_rectangle', rights')}
 | 
|  |   2260 |           end
 | 
|  |   2261 |      else                     (* column 0 is all constructors *)
 | 
|  |   2262 |      let val Type (ty_name,_) = type_of p
 | 
|  |   2263 |      in
 | 
|  |   2264 |      case (ty_info ty_name)
 | 
|  |   2265 |      of NONE => fail("Not a known datatype: "^ty_name)
 | 
|  |   2266 |       | SOME{constructors,nchotomy} =>
 | 
|  |   2267 |         let val thm' = Rules.ISPEC (tych u) nchotomy
 | 
|  |   2268 |             val disjuncts = USyntax.strip_disj (concl thm')
 | 
|  |   2269 |             val subproblems = divide(constructors, rows)
 | 
|  |   2270 |             val groups      = map #group subproblems
 | 
|  |   2271 |             and new_formals = map #new_formals subproblems
 | 
|  |   2272 |             val existentials = ListPair.map alpha_ex_unroll
 | 
|  |   2273 |                                    (new_formals, disjuncts)
 | 
|  |   2274 |             val constraints = map #1 existentials
 | 
|  |   2275 |             val vexl = map #2 existentials
 | 
|  |   2276 |             fun expnd tm (pats,(th,b)) = (pats, (Rules.SUBS ctxt [Rules.ASSUME (tych tm)] th, b))
 | 
|  |   2277 |             val news = map (fn (nf,rows,c) => {path = nf@rstp,
 | 
|  |   2278 |                                                rows = map (expnd c) rows})
 | 
|  |   2279 |                            (Utils.zip3 new_formals groups constraints)
 | 
|  |   2280 |             val recursive_thms = map mk news
 | 
|  |   2281 |             val build_exists = Library.foldr
 | 
|  |   2282 |                                 (fn((x,t), th) =>
 | 
|  |   2283 |                                  Rules.CHOOSE ctxt (tych x, Rules.ASSUME (tych t)) th)
 | 
|  |   2284 |             val thms' = ListPair.map build_exists (vexl, recursive_thms)
 | 
|  |   2285 |             val same_concls = Rules.EVEN_ORS thms'
 | 
|  |   2286 |         in Rules.DISJ_CASESL thm' same_concls
 | 
|  |   2287 |         end
 | 
|  |   2288 |      end end
 | 
|  |   2289 |  in mk
 | 
|  |   2290 |  end;
 | 
|  |   2291 | 
 | 
|  |   2292 | 
 | 
|  |   2293 | fun complete_cases thy =
 | 
|  |   2294 |  let val ctxt = Proof_Context.init_global thy
 | 
|  |   2295 |      val tych = Thry.typecheck thy
 | 
|  |   2296 |      val ty_info = Thry.induct_info thy
 | 
|  |   2297 |  in fn pats =>
 | 
|  |   2298 |  let val names = List.foldr Misc_Legacy.add_term_names [] pats
 | 
|  |   2299 |      val T = type_of (hd pats)
 | 
|  |   2300 |      val aname = singleton (Name.variant_list names) "a"
 | 
|  |   2301 |      val vname = singleton (Name.variant_list (aname::names)) "v"
 | 
|  |   2302 |      val a = Free (aname, T)
 | 
|  |   2303 |      val v = Free (vname, T)
 | 
|  |   2304 |      val a_eq_v = HOLogic.mk_eq(a,v)
 | 
|  |   2305 |      val ex_th0 = Rules.EXISTS (tych (USyntax.mk_exists{Bvar=v,Body=a_eq_v}), tych a)
 | 
|  |   2306 |                            (Rules.REFL (tych a))
 | 
|  |   2307 |      val th0 = Rules.ASSUME (tych a_eq_v)
 | 
|  |   2308 |      val rows = map (fn x => ([x], (th0,[]))) pats
 | 
|  |   2309 |  in
 | 
|  |   2310 |  Rules.GEN ctxt (tych a)
 | 
|  |   2311 |        (Rules.RIGHT_ASSOC ctxt
 | 
|  |   2312 |           (Rules.CHOOSE ctxt (tych v, ex_th0)
 | 
|  |   2313 |                 (mk_case ty_info (vname::aname::names)
 | 
|  |   2314 |                  thy {path=[v], rows=rows})))
 | 
|  |   2315 |  end end;
 | 
|  |   2316 | 
 | 
|  |   2317 | 
 | 
|  |   2318 | (*---------------------------------------------------------------------------
 | 
|  |   2319 |  * Constructing induction hypotheses: one for each recursive call.
 | 
|  |   2320 |  *
 | 
|  |   2321 |  * Note. R will never occur as a variable in the ind_clause, because
 | 
|  |   2322 |  * to do so, it would have to be from a nested definition, and we don't
 | 
|  |   2323 |  * allow nested defns to have R variable.
 | 
|  |   2324 |  *
 | 
|  |   2325 |  * Note. When the context is empty, there can be no local variables.
 | 
|  |   2326 |  *---------------------------------------------------------------------------*)
 | 
|  |   2327 | 
 | 
|  |   2328 | local infix 5 ==>
 | 
|  |   2329 |       fun (tm1 ==> tm2) = USyntax.mk_imp{ant = tm1, conseq = tm2}
 | 
|  |   2330 | in
 | 
|  |   2331 | fun build_ih f (P,SV) (pat,TCs) =
 | 
|  |   2332 |  let val pat_vars = USyntax.free_vars_lr pat
 | 
|  |   2333 |      val globals = pat_vars@SV
 | 
|  |   2334 |      fun nested tm = is_some (USyntax.find_term (curry (op aconv) f) tm)
 | 
|  |   2335 |      fun dest_TC tm =
 | 
|  |   2336 |          let val (cntxt,R_y_pat) = USyntax.strip_imp(#2(USyntax.strip_forall tm))
 | 
|  |   2337 |              val (R,y,_) = USyntax.dest_relation R_y_pat
 | 
|  |   2338 |              val P_y = if (nested tm) then R_y_pat ==> P$y else P$y
 | 
|  |   2339 |          in case cntxt
 | 
|  |   2340 |               of [] => (P_y, (tm,[]))
 | 
|  |   2341 |                | _  => let
 | 
|  |   2342 |                     val imp = USyntax.list_mk_conj cntxt ==> P_y
 | 
|  |   2343 |                     val lvs = subtract (op aconv) globals (USyntax.free_vars_lr imp)
 | 
|  |   2344 |                     val locals = #2(Utils.pluck (curry (op aconv) P) lvs) handle Utils.ERR _ => lvs
 | 
|  |   2345 |                     in (USyntax.list_mk_forall(locals,imp), (tm,locals)) end
 | 
|  |   2346 |          end
 | 
|  |   2347 |  in case TCs
 | 
|  |   2348 |     of [] => (USyntax.list_mk_forall(pat_vars, P$pat), [])
 | 
|  |   2349 |      |  _ => let val (ihs, TCs_locals) = ListPair.unzip(map dest_TC TCs)
 | 
|  |   2350 |                  val ind_clause = USyntax.list_mk_conj ihs ==> P$pat
 | 
|  |   2351 |              in (USyntax.list_mk_forall(pat_vars,ind_clause), TCs_locals)
 | 
|  |   2352 |              end
 | 
|  |   2353 |  end
 | 
|  |   2354 | end;
 | 
|  |   2355 | 
 | 
|  |   2356 | (*---------------------------------------------------------------------------
 | 
|  |   2357 |  * This function makes good on the promise made in "build_ih".
 | 
|  |   2358 |  *
 | 
|  |   2359 |  * Input  is tm = "(!y. R y pat ==> P y) ==> P pat",
 | 
|  |   2360 |  *           TCs = TC_1[pat] ... TC_n[pat]
 | 
|  |   2361 |  *           thm = ih1 /\ ... /\ ih_n |- ih[pat]
 | 
|  |   2362 |  *---------------------------------------------------------------------------*)
 | 
|  |   2363 | fun prove_case ctxt f (tm,TCs_locals,thm) =
 | 
|  |   2364 |  let val tych = Thry.typecheck (Proof_Context.theory_of ctxt)
 | 
|  |   2365 |      val antc = tych(#ant(USyntax.dest_imp tm))
 | 
|  |   2366 |      val thm' = Rules.SPEC_ALL thm
 | 
|  |   2367 |      fun nested tm = is_some (USyntax.find_term (curry (op aconv) f) tm)
 | 
|  |   2368 |      fun get_cntxt TC = tych(#ant(USyntax.dest_imp(#2(USyntax.strip_forall(concl TC)))))
 | 
|  |   2369 |      fun mk_ih ((TC,locals),th2,nested) =
 | 
|  |   2370 |          Rules.GENL ctxt (map tych locals)
 | 
|  |   2371 |             (if nested then Rules.DISCH (get_cntxt TC) th2 handle Utils.ERR _ => th2
 | 
|  |   2372 |              else if USyntax.is_imp (concl TC) then Rules.IMP_TRANS TC th2
 | 
|  |   2373 |              else Rules.MP th2 TC)
 | 
|  |   2374 |  in
 | 
|  |   2375 |  Rules.DISCH antc
 | 
|  |   2376 |  (if USyntax.is_imp(concl thm') (* recursive calls in this clause *)
 | 
|  |   2377 |   then let val th1 = Rules.ASSUME antc
 | 
|  |   2378 |            val TCs = map #1 TCs_locals
 | 
|  |   2379 |            val ylist = map (#2 o USyntax.dest_relation o #2 o USyntax.strip_imp o
 | 
|  |   2380 |                             #2 o USyntax.strip_forall) TCs
 | 
|  |   2381 |            val TClist = map (fn(TC,lvs) => (Rules.SPEC_ALL(Rules.ASSUME(tych TC)),lvs))
 | 
|  |   2382 |                             TCs_locals
 | 
|  |   2383 |            val th2list = map (fn t => Rules.SPEC (tych t) th1) ylist
 | 
|  |   2384 |            val nlist = map nested TCs
 | 
|  |   2385 |            val triples = Utils.zip3 TClist th2list nlist
 | 
|  |   2386 |            val Pylist = map mk_ih triples
 | 
|  |   2387 |        in Rules.MP thm' (Rules.LIST_CONJ Pylist) end
 | 
|  |   2388 |   else thm')
 | 
|  |   2389 |  end;
 | 
|  |   2390 | 
 | 
|  |   2391 | 
 | 
|  |   2392 | (*---------------------------------------------------------------------------
 | 
|  |   2393 |  *
 | 
|  |   2394 |  *         x = (v1,...,vn)  |- M[x]
 | 
|  |   2395 |  *    ---------------------------------------------
 | 
|  |   2396 |  *      ?v1 ... vn. x = (v1,...,vn) |- M[x]
 | 
|  |   2397 |  *
 | 
|  |   2398 |  *---------------------------------------------------------------------------*)
 | 
|  |   2399 | fun LEFT_ABS_VSTRUCT ctxt tych thm =
 | 
|  |   2400 |   let fun CHOOSER v (tm,thm) =
 | 
|  |   2401 |         let val ex_tm = USyntax.mk_exists{Bvar=v,Body=tm}
 | 
|  |   2402 |         in (ex_tm, Rules.CHOOSE ctxt (tych v, Rules.ASSUME (tych ex_tm)) thm)
 | 
|  |   2403 |         end
 | 
|  |   2404 |       val [veq] = filter (can USyntax.dest_eq) (#1 (Rules.dest_thm thm))
 | 
|  |   2405 |       val {lhs,rhs} = USyntax.dest_eq veq
 | 
|  |   2406 |       val L = USyntax.free_vars_lr rhs
 | 
|  |   2407 |   in  #2 (fold_rev CHOOSER L (veq,thm))  end;
 | 
|  |   2408 | 
 | 
|  |   2409 | 
 | 
|  |   2410 | (*----------------------------------------------------------------------------
 | 
|  |   2411 |  * Input : f, R,  and  [(pat1,TCs1),..., (patn,TCsn)]
 | 
|  |   2412 |  *
 | 
| 60522 |   2413 |  * Instantiates tfl_wf_induct, getting Sinduct and then tries to prove
 | 
| 60520 |   2414 |  * recursion induction (Rinduct) by proving the antecedent of Sinduct from
 | 
|  |   2415 |  * the antecedent of Rinduct.
 | 
|  |   2416 |  *---------------------------------------------------------------------------*)
 | 
|  |   2417 | fun mk_induction thy {fconst, R, SV, pat_TCs_list} =
 | 
|  |   2418 | let val ctxt = Proof_Context.init_global thy
 | 
|  |   2419 |     val tych = Thry.typecheck thy
 | 
| 60522 |   2420 |     val Sinduction = Rules.UNDISCH (Rules.ISPEC (tych R) @{thm tfl_wf_induct})
 | 
| 60520 |   2421 |     val (pats,TCsl) = ListPair.unzip pat_TCs_list
 | 
|  |   2422 |     val case_thm = complete_cases thy pats
 | 
|  |   2423 |     val domain = (type_of o hd) pats
 | 
|  |   2424 |     val Pname = singleton (Name.variant_list (List.foldr (Library.foldr Misc_Legacy.add_term_names)
 | 
|  |   2425 |                               [] (pats::TCsl))) "P"
 | 
|  |   2426 |     val P = Free(Pname, domain --> HOLogic.boolT)
 | 
|  |   2427 |     val Sinduct = Rules.SPEC (tych P) Sinduction
 | 
|  |   2428 |     val Sinduct_assumf = USyntax.rand ((#ant o USyntax.dest_imp o concl) Sinduct)
 | 
|  |   2429 |     val Rassums_TCl' = map (build_ih fconst (P,SV)) pat_TCs_list
 | 
|  |   2430 |     val (Rassums,TCl') = ListPair.unzip Rassums_TCl'
 | 
|  |   2431 |     val Rinduct_assum = Rules.ASSUME (tych (USyntax.list_mk_conj Rassums))
 | 
|  |   2432 |     val cases = map (fn pat => Term.betapply (Sinduct_assumf, pat)) pats
 | 
|  |   2433 |     val tasks = Utils.zip3 cases TCl' (Rules.CONJUNCTS Rinduct_assum)
 | 
|  |   2434 |     val proved_cases = map (prove_case ctxt fconst) tasks
 | 
|  |   2435 |     val v =
 | 
|  |   2436 |       Free (singleton
 | 
|  |   2437 |         (Name.variant_list (List.foldr Misc_Legacy.add_term_names [] (map concl proved_cases))) "v",
 | 
|  |   2438 |           domain)
 | 
|  |   2439 |     val vtyped = tych v
 | 
|  |   2440 |     val substs = map (Rules.SYM o Rules.ASSUME o tych o (curry HOLogic.mk_eq v)) pats
 | 
|  |   2441 |     val proved_cases1 = ListPair.map (fn (th,th') => Rules.SUBS ctxt [th]th')
 | 
|  |   2442 |                           (substs, proved_cases)
 | 
|  |   2443 |     val abs_cases = map (LEFT_ABS_VSTRUCT ctxt tych) proved_cases1
 | 
|  |   2444 |     val dant = Rules.GEN ctxt vtyped (Rules.DISJ_CASESL (Rules.ISPEC vtyped case_thm) abs_cases)
 | 
|  |   2445 |     val dc = Rules.MP Sinduct dant
 | 
|  |   2446 |     val Parg_ty = type_of(#Bvar(USyntax.dest_forall(concl dc)))
 | 
|  |   2447 |     val vars = map (gvvariant[Pname]) (USyntax.strip_prod_type Parg_ty)
 | 
|  |   2448 |     val dc' = fold_rev (Rules.GEN ctxt o tych) vars
 | 
|  |   2449 |                        (Rules.SPEC (tych(USyntax.mk_vstruct Parg_ty vars)) dc)
 | 
|  |   2450 | in
 | 
|  |   2451 |    Rules.GEN ctxt (tych P) (Rules.DISCH (tych(concl Rinduct_assum)) dc')
 | 
|  |   2452 | end
 | 
|  |   2453 | handle Utils.ERR _ => raise TFL_ERR "mk_induction" "failed derivation";
 | 
|  |   2454 | 
 | 
|  |   2455 | 
 | 
|  |   2456 | 
 | 
|  |   2457 | 
 | 
|  |   2458 | (*---------------------------------------------------------------------------
 | 
|  |   2459 |  *
 | 
|  |   2460 |  *                        POST PROCESSING
 | 
|  |   2461 |  *
 | 
|  |   2462 |  *---------------------------------------------------------------------------*)
 | 
|  |   2463 | 
 | 
|  |   2464 | 
 | 
|  |   2465 | fun simplify_induction thy hth ind =
 | 
|  |   2466 |   let val tych = Thry.typecheck thy
 | 
|  |   2467 |       val (asl,_) = Rules.dest_thm ind
 | 
|  |   2468 |       val (_,tc_eq_tc') = Rules.dest_thm hth
 | 
|  |   2469 |       val tc = USyntax.lhs tc_eq_tc'
 | 
|  |   2470 |       fun loop [] = ind
 | 
|  |   2471 |         | loop (asm::rst) =
 | 
|  |   2472 |           if (can (Thry.match_term thy asm) tc)
 | 
|  |   2473 |           then Rules.UNDISCH
 | 
|  |   2474 |                  (Rules.MATCH_MP
 | 
| 60522 |   2475 |                      (Rules.MATCH_MP @{thm tfl_simp_thm} (Rules.DISCH (tych asm) ind))
 | 
| 60520 |   2476 |                      hth)
 | 
|  |   2477 |          else loop rst
 | 
|  |   2478 |   in loop asl
 | 
|  |   2479 | end;
 | 
|  |   2480 | 
 | 
|  |   2481 | 
 | 
|  |   2482 | (*---------------------------------------------------------------------------
 | 
|  |   2483 |  * The termination condition is an antecedent to the rule, and an
 | 
|  |   2484 |  * assumption to the theorem.
 | 
|  |   2485 |  *---------------------------------------------------------------------------*)
 | 
|  |   2486 | fun elim_tc tcthm (rule,induction) =
 | 
|  |   2487 |    (Rules.MP rule tcthm, Rules.PROVE_HYP tcthm induction)
 | 
|  |   2488 | 
 | 
|  |   2489 | 
 | 
|  |   2490 | fun trace_thms ctxt s L =
 | 
|  |   2491 |   if !trace then writeln (cat_lines (s :: map (Display.string_of_thm ctxt) L))
 | 
|  |   2492 |   else ();
 | 
|  |   2493 | 
 | 
|  |   2494 | fun trace_cterm ctxt s ct =
 | 
|  |   2495 |   if !trace then
 | 
|  |   2496 |     writeln (cat_lines [s, Syntax.string_of_term ctxt (Thm.term_of ct)])
 | 
|  |   2497 |   else ();
 | 
|  |   2498 | 
 | 
|  |   2499 | 
 | 
|  |   2500 | fun postprocess ctxt strict {wf_tac, terminator, simplifier} {rules,induction,TCs} =
 | 
|  |   2501 |   let
 | 
|  |   2502 |     val thy = Proof_Context.theory_of ctxt;
 | 
|  |   2503 |     val tych = Thry.typecheck thy;
 | 
|  |   2504 | 
 | 
|  |   2505 |    (*---------------------------------------------------------------------
 | 
|  |   2506 |     * Attempt to eliminate WF condition. It's the only assumption of rules
 | 
|  |   2507 |     *---------------------------------------------------------------------*)
 | 
|  |   2508 |     val (rules1,induction1)  =
 | 
|  |   2509 |        let val thm =
 | 
|  |   2510 |         Rules.prove ctxt strict (HOLogic.mk_Trueprop (hd(#1(Rules.dest_thm rules))), wf_tac)
 | 
|  |   2511 |        in (Rules.PROVE_HYP thm rules, Rules.PROVE_HYP thm induction)
 | 
|  |   2512 |        end handle Utils.ERR _ => (rules,induction);
 | 
|  |   2513 | 
 | 
|  |   2514 |    (*----------------------------------------------------------------------
 | 
|  |   2515 |     * The termination condition (tc) is simplified to |- tc = tc' (there
 | 
|  |   2516 |     * might not be a change!) and then 3 attempts are made:
 | 
|  |   2517 |     *
 | 
| 60522 |   2518 |     *   1. if |- tc = T, then eliminate it with tfl_eq_True; otherwise,
 | 
| 60520 |   2519 |     *   2. apply the terminator to tc'. If |- tc' = T then eliminate; else
 | 
|  |   2520 |     *   3. replace tc by tc' in both the rules and the induction theorem.
 | 
|  |   2521 |     *---------------------------------------------------------------------*)
 | 
|  |   2522 | 
 | 
|  |   2523 |    fun simplify_tc tc (r,ind) =
 | 
|  |   2524 |        let val tc1 = tych tc
 | 
|  |   2525 |            val _ = trace_cterm ctxt "TC before simplification: " tc1
 | 
|  |   2526 |            val tc_eq = simplifier tc1
 | 
|  |   2527 |            val _ = trace_thms ctxt "result: " [tc_eq]
 | 
|  |   2528 |        in
 | 
| 60522 |   2529 |        elim_tc (Rules.MATCH_MP @{thm tfl_eq_True} tc_eq) (r,ind)
 | 
| 60520 |   2530 |        handle Utils.ERR _ =>
 | 
| 60522 |   2531 |         (elim_tc (Rules.MATCH_MP(Rules.MATCH_MP @{thm tfl_rev_eq_mp} tc_eq)
 | 
| 60520 |   2532 |                   (Rules.prove ctxt strict (HOLogic.mk_Trueprop(USyntax.rhs(concl tc_eq)),
 | 
|  |   2533 |                            terminator)))
 | 
|  |   2534 |                  (r,ind)
 | 
|  |   2535 |          handle Utils.ERR _ =>
 | 
| 60522 |   2536 |           (Rules.UNDISCH(Rules.MATCH_MP (Rules.MATCH_MP @{thm tfl_simp_thm} r) tc_eq),
 | 
| 60520 |   2537 |            simplify_induction thy tc_eq ind))
 | 
|  |   2538 |        end
 | 
|  |   2539 | 
 | 
|  |   2540 |    (*----------------------------------------------------------------------
 | 
|  |   2541 |     * Nested termination conditions are harder to get at, since they are
 | 
|  |   2542 |     * left embedded in the body of the function (and in induction
 | 
|  |   2543 |     * theorem hypotheses). Our "solution" is to simplify them, and try to
 | 
|  |   2544 |     * prove termination, but leave the application of the resulting theorem
 | 
|  |   2545 |     * to a higher level. So things go much as in "simplify_tc": the
 | 
|  |   2546 |     * termination condition (tc) is simplified to |- tc = tc' (there might
 | 
|  |   2547 |     * not be a change) and then 2 attempts are made:
 | 
|  |   2548 |     *
 | 
|  |   2549 |     *   1. if |- tc = T, then return |- tc; otherwise,
 | 
|  |   2550 |     *   2. apply the terminator to tc'. If |- tc' = T then return |- tc; else
 | 
|  |   2551 |     *   3. return |- tc = tc'
 | 
|  |   2552 |     *---------------------------------------------------------------------*)
 | 
|  |   2553 |    fun simplify_nested_tc tc =
 | 
|  |   2554 |       let val tc_eq = simplifier (tych (#2 (USyntax.strip_forall tc)))
 | 
|  |   2555 |       in
 | 
|  |   2556 |       Rules.GEN_ALL ctxt
 | 
| 60522 |   2557 |        (Rules.MATCH_MP @{thm tfl_eq_True} tc_eq
 | 
| 60520 |   2558 |         handle Utils.ERR _ =>
 | 
| 60522 |   2559 |           (Rules.MATCH_MP(Rules.MATCH_MP @{thm tfl_rev_eq_mp} tc_eq)
 | 
| 60520 |   2560 |                       (Rules.prove ctxt strict (HOLogic.mk_Trueprop (USyntax.rhs(concl tc_eq)),
 | 
|  |   2561 |                                terminator))
 | 
|  |   2562 |             handle Utils.ERR _ => tc_eq))
 | 
|  |   2563 |       end
 | 
|  |   2564 | 
 | 
|  |   2565 |    (*-------------------------------------------------------------------
 | 
|  |   2566 |     * Attempt to simplify the termination conditions in each rule and
 | 
|  |   2567 |     * in the induction theorem.
 | 
|  |   2568 |     *-------------------------------------------------------------------*)
 | 
|  |   2569 |    fun strip_imp tm = if USyntax.is_neg tm then ([],tm) else USyntax.strip_imp tm
 | 
|  |   2570 |    fun loop ([],extras,R,ind) = (rev R, ind, extras)
 | 
|  |   2571 |      | loop ((r,ftcs)::rst, nthms, R, ind) =
 | 
|  |   2572 |         let val tcs = #1(strip_imp (concl r))
 | 
|  |   2573 |             val extra_tcs = subtract (op aconv) tcs ftcs
 | 
|  |   2574 |             val extra_tc_thms = map simplify_nested_tc extra_tcs
 | 
|  |   2575 |             val (r1,ind1) = fold simplify_tc tcs (r,ind)
 | 
|  |   2576 |             val r2 = Rules.FILTER_DISCH_ALL(not o USyntax.is_WFR) r1
 | 
|  |   2577 |         in loop(rst, nthms@extra_tc_thms, r2::R, ind1)
 | 
|  |   2578 |         end
 | 
|  |   2579 |    val rules_tcs = ListPair.zip (Rules.CONJUNCTS rules1, TCs)
 | 
|  |   2580 |    val (rules2,ind2,extras) = loop(rules_tcs,[],[],induction1)
 | 
|  |   2581 | in
 | 
|  |   2582 |   {induction = ind2, rules = Rules.LIST_CONJ rules2, nested_tcs = extras}
 | 
|  |   2583 | end;
 | 
|  |   2584 | 
 | 
|  |   2585 | end;
 | 
|  |   2586 | 
 | 
|  |   2587 | 
 | 
| 60521 |   2588 | 
 | 
| 60520 |   2589 | (*** second part of main module (postprocessing of TFL definitions) ***)
 | 
|  |   2590 | 
 | 
|  |   2591 | structure Tfl: TFL =
 | 
|  |   2592 | struct
 | 
|  |   2593 | 
 | 
|  |   2594 | (* misc *)
 | 
|  |   2595 | 
 | 
|  |   2596 | (*---------------------------------------------------------------------------
 | 
|  |   2597 |  * Extract termination goals so that they can be put it into a goalstack, or
 | 
|  |   2598 |  * have a tactic directly applied to them.
 | 
|  |   2599 |  *--------------------------------------------------------------------------*)
 | 
|  |   2600 | fun termination_goals rules =
 | 
|  |   2601 |     map (Type.legacy_freeze o HOLogic.dest_Trueprop)
 | 
|  |   2602 |       (fold_rev (union (op aconv) o Thm.prems_of) rules []);
 | 
|  |   2603 | 
 | 
|  |   2604 | (*---------------------------------------------------------------------------
 | 
|  |   2605 |  * Three postprocessors are applied to the definition.  It
 | 
|  |   2606 |  * attempts to prove wellfoundedness of the given relation, simplifies the
 | 
|  |   2607 |  * non-proved termination conditions, and finally attempts to prove the
 | 
|  |   2608 |  * simplified termination conditions.
 | 
|  |   2609 |  *--------------------------------------------------------------------------*)
 | 
|  |   2610 | fun std_postprocessor ctxt strict wfs =
 | 
|  |   2611 |   Prim.postprocess ctxt strict
 | 
|  |   2612 |    {wf_tac = REPEAT (ares_tac wfs 1),
 | 
|  |   2613 |     terminator =
 | 
|  |   2614 |       asm_simp_tac ctxt 1
 | 
|  |   2615 |       THEN TRY (Arith_Data.arith_tac ctxt 1 ORELSE
 | 
|  |   2616 |         fast_force_tac (ctxt addSDs @{thms not0_implies_Suc}) 1),
 | 
|  |   2617 |     simplifier = Rules.simpl_conv ctxt []};
 | 
|  |   2618 | 
 | 
|  |   2619 | 
 | 
|  |   2620 | 
 | 
|  |   2621 | val concl = #2 o Rules.dest_thm;
 | 
|  |   2622 | 
 | 
|  |   2623 | (*---------------------------------------------------------------------------
 | 
|  |   2624 |  * Postprocess a definition made by "define". This is a separate stage of
 | 
|  |   2625 |  * processing from the definition stage.
 | 
|  |   2626 |  *---------------------------------------------------------------------------*)
 | 
|  |   2627 | local
 | 
|  |   2628 | 
 | 
|  |   2629 | (* The rest of these local definitions are for the tricky nested case *)
 | 
|  |   2630 | val solved = not o can USyntax.dest_eq o #2 o USyntax.strip_forall o concl
 | 
|  |   2631 | 
 | 
|  |   2632 | fun id_thm th =
 | 
|  |   2633 |    let val {lhs,rhs} = USyntax.dest_eq (#2 (USyntax.strip_forall (#2 (Rules.dest_thm th))));
 | 
|  |   2634 |    in lhs aconv rhs end
 | 
|  |   2635 |    handle Utils.ERR _ => false;
 | 
|  |   2636 | 
 | 
|  |   2637 | val P_imp_P_eq_True = @{thm eqTrueI} RS eq_reflection;
 | 
|  |   2638 | fun mk_meta_eq r =
 | 
|  |   2639 |   (case Thm.concl_of r of
 | 
|  |   2640 |      Const(@{const_name Pure.eq},_)$_$_ => r
 | 
|  |   2641 |   |   _ $(Const(@{const_name HOL.eq},_)$_$_) => r RS eq_reflection
 | 
|  |   2642 |   |   _ => r RS P_imp_P_eq_True)
 | 
|  |   2643 | 
 | 
|  |   2644 | (*Is this the best way to invoke the simplifier??*)
 | 
|  |   2645 | fun rewrite ctxt L = rewrite_rule ctxt (map mk_meta_eq (filter_out id_thm L))
 | 
|  |   2646 | 
 | 
|  |   2647 | fun join_assums ctxt th =
 | 
|  |   2648 |   let val tych = Thm.cterm_of ctxt
 | 
|  |   2649 |       val {lhs,rhs} = USyntax.dest_eq(#2 (USyntax.strip_forall (concl th)))
 | 
|  |   2650 |       val cntxtl = (#1 o USyntax.strip_imp) lhs  (* cntxtl should = cntxtr *)
 | 
|  |   2651 |       val cntxtr = (#1 o USyntax.strip_imp) rhs  (* but union is solider *)
 | 
|  |   2652 |       val cntxt = union (op aconv) cntxtl cntxtr
 | 
|  |   2653 |   in
 | 
|  |   2654 |     Rules.GEN_ALL ctxt
 | 
|  |   2655 |       (Rules.DISCH_ALL
 | 
|  |   2656 |          (rewrite ctxt (map (Rules.ASSUME o tych) cntxt) (Rules.SPEC_ALL th)))
 | 
|  |   2657 |   end
 | 
|  |   2658 |   val gen_all = USyntax.gen_all
 | 
|  |   2659 | in
 | 
|  |   2660 | fun proof_stage ctxt strict wfs {f, R, rules, full_pats_TCs, TCs} =
 | 
|  |   2661 |   let
 | 
|  |   2662 |     val _ = writeln "Proving induction theorem ..."
 | 
|  |   2663 |     val ind =
 | 
|  |   2664 |       Prim.mk_induction (Proof_Context.theory_of ctxt)
 | 
|  |   2665 |         {fconst=f, R=R, SV=[], pat_TCs_list=full_pats_TCs}
 | 
|  |   2666 |     val _ = writeln "Postprocessing ...";
 | 
|  |   2667 |     val {rules, induction, nested_tcs} =
 | 
|  |   2668 |       std_postprocessor ctxt strict wfs {rules=rules, induction=ind, TCs=TCs}
 | 
|  |   2669 |   in
 | 
|  |   2670 |   case nested_tcs
 | 
|  |   2671 |   of [] => {induction=induction, rules=rules,tcs=[]}
 | 
| 60521 |   2672 |   | L  => let val _ = writeln "Simplifying nested TCs ..."
 | 
| 60520 |   2673 |               val (solved,simplified,stubborn) =
 | 
|  |   2674 |                fold_rev (fn th => fn (So,Si,St) =>
 | 
|  |   2675 |                      if (id_thm th) then (So, Si, th::St) else
 | 
|  |   2676 |                      if (solved th) then (th::So, Si, St)
 | 
|  |   2677 |                      else (So, th::Si, St)) nested_tcs ([],[],[])
 | 
|  |   2678 |               val simplified' = map (join_assums ctxt) simplified
 | 
|  |   2679 |               val dummy = (Prim.trace_thms ctxt "solved =" solved;
 | 
|  |   2680 |                            Prim.trace_thms ctxt "simplified' =" simplified')
 | 
|  |   2681 |               val rewr = full_simplify (ctxt addsimps (solved @ simplified'));
 | 
|  |   2682 |               val dummy = Prim.trace_thms ctxt "Simplifying the induction rule..." [induction]
 | 
|  |   2683 |               val induction' = rewr induction
 | 
|  |   2684 |               val dummy = Prim.trace_thms ctxt "Simplifying the recursion rules..." [rules]
 | 
|  |   2685 |               val rules'     = rewr rules
 | 
|  |   2686 |               val _ = writeln "... Postprocessing finished";
 | 
|  |   2687 |           in
 | 
|  |   2688 |           {induction = induction',
 | 
|  |   2689 |                rules = rules',
 | 
|  |   2690 |                  tcs = map (gen_all o USyntax.rhs o #2 o USyntax.strip_forall o concl)
 | 
|  |   2691 |                            (simplified@stubborn)}
 | 
|  |   2692 |           end
 | 
|  |   2693 |   end;
 | 
|  |   2694 | 
 | 
|  |   2695 | 
 | 
|  |   2696 | (*lcp: curry the predicate of the induction rule*)
 | 
|  |   2697 | fun curry_rule ctxt rl =
 | 
|  |   2698 |   Split_Rule.split_rule_var ctxt (Term.head_of (HOLogic.dest_Trueprop (Thm.concl_of rl))) rl;
 | 
|  |   2699 | 
 | 
|  |   2700 | (*lcp: put a theorem into Isabelle form, using meta-level connectives*)
 | 
|  |   2701 | fun meta_outer ctxt =
 | 
|  |   2702 |   curry_rule ctxt o Drule.export_without_context o
 | 
|  |   2703 |   rule_by_tactic ctxt (REPEAT (FIRSTGOAL (resolve_tac ctxt [allI, impI, conjI] ORELSE' etac conjE)));
 | 
|  |   2704 | 
 | 
|  |   2705 | (*Strip off the outer !P*)
 | 
|  |   2706 | val spec'=
 | 
|  |   2707 |   Rule_Insts.read_instantiate @{context} [((("x", 0), Position.none), "P::'b=>bool")] [] spec;
 | 
|  |   2708 | 
 | 
| 60524 |   2709 | fun simplify_defn ctxt strict congs wfs pats def0 =
 | 
| 60520 |   2710 |   let
 | 
|  |   2711 |     val def = Thm.unvarify_global def0 RS meta_eq_to_obj_eq
 | 
|  |   2712 |     val {rules, rows, TCs, full_pats_TCs} = Prim.post_definition ctxt congs (def, pats)
 | 
|  |   2713 |     val {lhs=f,rhs} = USyntax.dest_eq (concl def)
 | 
|  |   2714 |     val (_,[R,_]) = USyntax.strip_comb rhs
 | 
| 60521 |   2715 |     val _ = Prim.trace_thms ctxt "congs =" congs
 | 
| 60520 |   2716 |     (*the next step has caused simplifier looping in some cases*)
 | 
|  |   2717 |     val {induction, rules, tcs} =
 | 
|  |   2718 |       proof_stage ctxt strict wfs
 | 
|  |   2719 |        {f = f, R = R, rules = rules,
 | 
|  |   2720 |         full_pats_TCs = full_pats_TCs,
 | 
|  |   2721 |         TCs = TCs}
 | 
|  |   2722 |     val rules' = map (Drule.export_without_context o Object_Logic.rulify_no_asm ctxt)
 | 
|  |   2723 |                       (Rules.CONJUNCTS rules)
 | 
|  |   2724 |   in
 | 
|  |   2725 |     {induct = meta_outer ctxt (Object_Logic.rulify_no_asm ctxt (induction RS spec')),
 | 
|  |   2726 |      rules = ListPair.zip(rules', rows),
 | 
|  |   2727 |      tcs = (termination_goals rules') @ tcs}
 | 
|  |   2728 |   end
 | 
|  |   2729 |   handle Utils.ERR {mesg,func,module} =>
 | 
|  |   2730 |     error (mesg ^ "\n    (In TFL function " ^ module ^ "." ^ func ^ ")");
 | 
|  |   2731 | 
 | 
|  |   2732 | 
 | 
|  |   2733 | (* Derive the initial equations from the case-split rules to meet the
 | 
|  |   2734 | users specification of the recursive function. *)
 | 
|  |   2735 | local
 | 
|  |   2736 |   fun get_related_thms i =
 | 
|  |   2737 |       map_filter ((fn (r,x) => if x = i then SOME r else NONE));
 | 
|  |   2738 | 
 | 
| 60521 |   2739 |   fun solve_eq _ (_, [], _) =  error "derive_init_eqs: missing rules"
 | 
|  |   2740 |     | solve_eq _ (_, [a], i) = [(a, i)]
 | 
| 60520 |   2741 |     | solve_eq ctxt (th, splitths, i) =
 | 
|  |   2742 |       (writeln "Proving unsplit equation...";
 | 
|  |   2743 |       [((Drule.export_without_context o Object_Logic.rulify_no_asm ctxt)
 | 
|  |   2744 |           (CaseSplit.splitto ctxt splitths th), i)])
 | 
|  |   2745 |       handle ERROR s =>
 | 
|  |   2746 |              (warning ("recdef (solve_eq): " ^ s); map (fn x => (x,i)) splitths);
 | 
|  |   2747 | in
 | 
|  |   2748 | fun derive_init_eqs ctxt rules eqs =
 | 
|  |   2749 |   map (Thm.trivial o Thm.cterm_of ctxt o HOLogic.mk_Trueprop) eqs
 | 
|  |   2750 |   |> map_index (fn (i, e) => solve_eq ctxt (e, (get_related_thms i rules), i))
 | 
|  |   2751 |   |> flat;
 | 
|  |   2752 | end;
 | 
|  |   2753 | 
 | 
|  |   2754 | 
 | 
|  |   2755 | (*---------------------------------------------------------------------------
 | 
|  |   2756 |  * Defining a function with an associated termination relation.
 | 
|  |   2757 |  *---------------------------------------------------------------------------*)
 | 
|  |   2758 | fun define_i strict congs wfs fid R eqs ctxt =
 | 
|  |   2759 |   let
 | 
|  |   2760 |     val thy = Proof_Context.theory_of ctxt
 | 
|  |   2761 |     val {functional, pats} = Prim.mk_functional thy eqs
 | 
|  |   2762 |     val (def, thy') = Prim.wfrec_definition0 fid R functional thy
 | 
|  |   2763 |     val ctxt' = Proof_Context.transfer thy' ctxt
 | 
|  |   2764 |     val (lhs, _) = Logic.dest_equals (Thm.prop_of def)
 | 
| 60524 |   2765 |     val {induct, rules, tcs} = simplify_defn ctxt' strict congs wfs pats def
 | 
| 60520 |   2766 |     val rules' = if strict then derive_init_eqs ctxt' rules eqs else rules
 | 
|  |   2767 |   in ({lhs = lhs, rules = rules', induct = induct, tcs = tcs}, ctxt') end;
 | 
|  |   2768 | 
 | 
|  |   2769 | fun define strict congs wfs fid R seqs ctxt =
 | 
|  |   2770 |   define_i strict congs wfs fid
 | 
|  |   2771 |     (Syntax.read_term ctxt R) (map (Syntax.read_term ctxt) seqs) ctxt
 | 
|  |   2772 |       handle Utils.ERR {mesg,...} => error mesg;
 | 
|  |   2773 | 
 | 
|  |   2774 | end;
 | 
|  |   2775 | 
 | 
|  |   2776 | end;
 | 
|  |   2777 | 
 | 
|  |   2778 | 
 | 
| 60521 |   2779 | 
 | 
| 60520 |   2780 | (*** wrappers for Isar ***)
 | 
|  |   2781 | 
 | 
|  |   2782 | (** recdef hints **)
 | 
|  |   2783 | 
 | 
|  |   2784 | (* type hints *)
 | 
|  |   2785 | 
 | 
|  |   2786 | type hints = {simps: thm list, congs: (string * thm) list, wfs: thm list};
 | 
|  |   2787 | 
 | 
|  |   2788 | fun mk_hints (simps, congs, wfs) = {simps = simps, congs = congs, wfs = wfs}: hints;
 | 
|  |   2789 | fun map_hints f ({simps, congs, wfs}: hints) = mk_hints (f (simps, congs, wfs));
 | 
|  |   2790 | 
 | 
|  |   2791 | fun map_simps f = map_hints (fn (simps, congs, wfs) => (f simps, congs, wfs));
 | 
|  |   2792 | fun map_congs f = map_hints (fn (simps, congs, wfs) => (simps, f congs, wfs));
 | 
|  |   2793 | fun map_wfs f = map_hints (fn (simps, congs, wfs) => (simps, congs, f wfs));
 | 
|  |   2794 | 
 | 
|  |   2795 | 
 | 
|  |   2796 | (* congruence rules *)
 | 
|  |   2797 | 
 | 
|  |   2798 | local
 | 
|  |   2799 | 
 | 
|  |   2800 | val cong_head =
 | 
|  |   2801 |   fst o Term.dest_Const o Term.head_of o fst o Logic.dest_equals o Thm.concl_of;
 | 
|  |   2802 | 
 | 
|  |   2803 | fun prep_cong raw_thm =
 | 
|  |   2804 |   let val thm = safe_mk_meta_eq raw_thm in (cong_head thm, thm) end;
 | 
|  |   2805 | 
 | 
|  |   2806 | in
 | 
|  |   2807 | 
 | 
|  |   2808 | fun add_cong raw_thm congs =
 | 
|  |   2809 |   let
 | 
|  |   2810 |     val (c, thm) = prep_cong raw_thm;
 | 
|  |   2811 |     val _ = if AList.defined (op =) congs c
 | 
|  |   2812 |       then warning ("Overwriting recdef congruence rule for " ^ quote c)
 | 
|  |   2813 |       else ();
 | 
|  |   2814 |   in AList.update (op =) (c, thm) congs end;
 | 
|  |   2815 | 
 | 
|  |   2816 | fun del_cong raw_thm congs =
 | 
|  |   2817 |   let
 | 
| 60524 |   2818 |     val (c, _) = prep_cong raw_thm;
 | 
| 60520 |   2819 |     val _ = if AList.defined (op =) congs c
 | 
|  |   2820 |       then ()
 | 
|  |   2821 |       else warning ("No recdef congruence rule for " ^ quote c);
 | 
|  |   2822 |   in AList.delete (op =) c congs end;
 | 
|  |   2823 | 
 | 
|  |   2824 | end;
 | 
|  |   2825 | 
 | 
|  |   2826 | 
 | 
|  |   2827 | 
 | 
|  |   2828 | (** global and local recdef data **)
 | 
|  |   2829 | 
 | 
|  |   2830 | (* theory data *)
 | 
|  |   2831 | 
 | 
|  |   2832 | type recdef_info = {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list};
 | 
|  |   2833 | 
 | 
|  |   2834 | structure Data = Generic_Data
 | 
|  |   2835 | (
 | 
|  |   2836 |   type T = recdef_info Symtab.table * hints;
 | 
|  |   2837 |   val empty = (Symtab.empty, mk_hints ([], [], [])): T;
 | 
|  |   2838 |   val extend = I;
 | 
|  |   2839 |   fun merge
 | 
|  |   2840 |    ((tab1, {simps = simps1, congs = congs1, wfs = wfs1}),
 | 
|  |   2841 |     (tab2, {simps = simps2, congs = congs2, wfs = wfs2})) : T =
 | 
|  |   2842 |       (Symtab.merge (K true) (tab1, tab2),
 | 
|  |   2843 |         mk_hints (Thm.merge_thms (simps1, simps2),
 | 
|  |   2844 |           AList.merge (op =) (K true) (congs1, congs2),
 | 
|  |   2845 |           Thm.merge_thms (wfs1, wfs2)));
 | 
|  |   2846 | );
 | 
|  |   2847 | 
 | 
|  |   2848 | val get_recdef = Symtab.lookup o #1 o Data.get o Context.Theory;
 | 
|  |   2849 | 
 | 
|  |   2850 | fun put_recdef name info =
 | 
|  |   2851 |   (Context.theory_map o Data.map o apfst) (fn tab =>
 | 
|  |   2852 |     Symtab.update_new (name, info) tab
 | 
|  |   2853 |       handle Symtab.DUP _ => error ("Duplicate recursive function definition " ^ quote name));
 | 
|  |   2854 | 
 | 
|  |   2855 | val get_hints = #2 o Data.get o Context.Proof;
 | 
|  |   2856 | val map_hints = Data.map o apsnd;
 | 
|  |   2857 | 
 | 
|  |   2858 | 
 | 
|  |   2859 | (* attributes *)
 | 
|  |   2860 | 
 | 
|  |   2861 | fun attrib f = Thm.declaration_attribute (map_hints o f);
 | 
|  |   2862 | 
 | 
|  |   2863 | val simp_add = attrib (map_simps o Thm.add_thm);
 | 
|  |   2864 | val simp_del = attrib (map_simps o Thm.del_thm);
 | 
|  |   2865 | val cong_add = attrib (map_congs o add_cong);
 | 
|  |   2866 | val cong_del = attrib (map_congs o del_cong);
 | 
|  |   2867 | val wf_add = attrib (map_wfs o Thm.add_thm);
 | 
|  |   2868 | val wf_del = attrib (map_wfs o Thm.del_thm);
 | 
|  |   2869 | 
 | 
|  |   2870 | 
 | 
|  |   2871 | (* modifiers *)
 | 
|  |   2872 | 
 | 
|  |   2873 | val recdef_simpN = "recdef_simp";
 | 
|  |   2874 | val recdef_congN = "recdef_cong";
 | 
|  |   2875 | val recdef_wfN = "recdef_wf";
 | 
|  |   2876 | 
 | 
|  |   2877 | val recdef_modifiers =
 | 
|  |   2878 |  [Args.$$$ recdef_simpN -- Args.colon >> K (Method.modifier simp_add @{here}),
 | 
|  |   2879 |   Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (Method.modifier simp_add @{here}),
 | 
|  |   2880 |   Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (Method.modifier simp_del @{here}),
 | 
|  |   2881 |   Args.$$$ recdef_congN -- Args.colon >> K (Method.modifier cong_add @{here}),
 | 
|  |   2882 |   Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (Method.modifier cong_add @{here}),
 | 
|  |   2883 |   Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (Method.modifier cong_del @{here}),
 | 
|  |   2884 |   Args.$$$ recdef_wfN -- Args.colon >> K (Method.modifier wf_add @{here}),
 | 
|  |   2885 |   Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (Method.modifier wf_add @{here}),
 | 
|  |   2886 |   Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (Method.modifier wf_del @{here})] @
 | 
|  |   2887 |   Clasimp.clasimp_modifiers;
 | 
|  |   2888 | 
 | 
|  |   2889 | 
 | 
|  |   2890 | 
 | 
|  |   2891 | (** prepare hints **)
 | 
|  |   2892 | 
 | 
|  |   2893 | fun prepare_hints opt_src ctxt =
 | 
|  |   2894 |   let
 | 
|  |   2895 |     val ctxt' =
 | 
|  |   2896 |       (case opt_src of
 | 
|  |   2897 |         NONE => ctxt
 | 
|  |   2898 |       | SOME src => #2 (Token.syntax (Method.sections recdef_modifiers) src ctxt));
 | 
|  |   2899 |     val {simps, congs, wfs} = get_hints ctxt';
 | 
|  |   2900 |     val ctxt'' = ctxt' addsimps simps |> Simplifier.del_cong @{thm imp_cong};
 | 
|  |   2901 |   in ((rev (map snd congs), wfs), ctxt'') end;
 | 
|  |   2902 | 
 | 
|  |   2903 | fun prepare_hints_i () ctxt =
 | 
|  |   2904 |   let
 | 
|  |   2905 |     val {simps, congs, wfs} = get_hints ctxt;
 | 
|  |   2906 |     val ctxt' = ctxt addsimps simps |> Simplifier.del_cong @{thm imp_cong};
 | 
|  |   2907 |   in ((rev (map snd congs), wfs), ctxt') end;
 | 
|  |   2908 | 
 | 
|  |   2909 | 
 | 
|  |   2910 | 
 | 
|  |   2911 | (** add_recdef(_i) **)
 | 
|  |   2912 | 
 | 
|  |   2913 | fun gen_add_recdef tfl_fn prep_att prep_hints not_permissive raw_name R eq_srcs hints thy =
 | 
|  |   2914 |   let
 | 
|  |   2915 |     val _ = legacy_feature "Old 'recdef' command -- use 'fun' or 'function' instead";
 | 
|  |   2916 | 
 | 
|  |   2917 |     val name = Sign.intern_const thy raw_name;
 | 
|  |   2918 |     val bname = Long_Name.base_name name;
 | 
|  |   2919 |     val _ = writeln ("Defining recursive function " ^ quote name ^ " ...");
 | 
|  |   2920 | 
 | 
|  |   2921 |     val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs);
 | 
|  |   2922 |     val eq_atts = map (map (prep_att thy)) raw_eq_atts;
 | 
|  |   2923 | 
 | 
|  |   2924 |     val ((congs, wfs), ctxt) = prep_hints hints (Proof_Context.init_global thy);
 | 
|  |   2925 |     (*We must remove imp_cong to prevent looping when the induction rule
 | 
|  |   2926 |       is simplified. Many induction rules have nested implications that would
 | 
|  |   2927 |       give rise to looping conditional rewriting.*)
 | 
|  |   2928 |     val ({lhs, rules = rules_idx, induct, tcs}, ctxt1) =
 | 
|  |   2929 |       tfl_fn not_permissive congs wfs name R eqs ctxt;
 | 
|  |   2930 |     val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx);
 | 
|  |   2931 |     val simp_att =
 | 
|  |   2932 |       if null tcs then [Simplifier.simp_add,
 | 
|  |   2933 |         Named_Theorems.add @{named_theorems nitpick_simp}, Code.add_default_eqn_attribute]
 | 
|  |   2934 |       else [];
 | 
|  |   2935 |     val ((simps' :: rules', [induct']), thy2) =
 | 
|  |   2936 |       Proof_Context.theory_of ctxt1
 | 
|  |   2937 |       |> Sign.add_path bname
 | 
|  |   2938 |       |> Global_Theory.add_thmss
 | 
|  |   2939 |         (((Binding.name "simps", flat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts))
 | 
|  |   2940 |       ||>> Global_Theory.add_thms [((Binding.name "induct", induct), [])]
 | 
|  |   2941 |       ||> Spec_Rules.add_global Spec_Rules.Equational ([lhs], flat rules);
 | 
|  |   2942 |     val result = {lhs = lhs, simps = simps', rules = rules', induct = induct', tcs = tcs};
 | 
|  |   2943 |     val thy3 =
 | 
|  |   2944 |       thy2
 | 
|  |   2945 |       |> put_recdef name result
 | 
|  |   2946 |       |> Sign.parent_path;
 | 
|  |   2947 |   in (thy3, result) end;
 | 
|  |   2948 | 
 | 
|  |   2949 | val add_recdef = gen_add_recdef Tfl.define Attrib.attribute_cmd_global prepare_hints;
 | 
|  |   2950 | fun add_recdef_i x y z w = gen_add_recdef Tfl.define_i (K I) prepare_hints_i x y z w ();
 | 
|  |   2951 | 
 | 
|  |   2952 | 
 | 
|  |   2953 | 
 | 
|  |   2954 | (** package setup **)
 | 
|  |   2955 | 
 | 
|  |   2956 | (* setup theory *)
 | 
|  |   2957 | 
 | 
|  |   2958 | val _ =
 | 
|  |   2959 |   Theory.setup
 | 
|  |   2960 |    (Attrib.setup @{binding recdef_simp} (Attrib.add_del simp_add simp_del)
 | 
|  |   2961 |       "declaration of recdef simp rule" #>
 | 
|  |   2962 |     Attrib.setup @{binding recdef_cong} (Attrib.add_del cong_add cong_del)
 | 
|  |   2963 |       "declaration of recdef cong rule" #>
 | 
|  |   2964 |     Attrib.setup @{binding recdef_wf} (Attrib.add_del wf_add wf_del)
 | 
|  |   2965 |       "declaration of recdef wf rule");
 | 
|  |   2966 | 
 | 
|  |   2967 | 
 | 
|  |   2968 | (* outer syntax *)
 | 
|  |   2969 | 
 | 
|  |   2970 | val hints =
 | 
|  |   2971 |   @{keyword "("} |--
 | 
|  |   2972 |     Parse.!!! (Parse.position @{keyword "hints"} -- Parse.args --| @{keyword ")"})
 | 
|  |   2973 |   >> uncurry Token.src;
 | 
|  |   2974 | 
 | 
|  |   2975 | val recdef_decl =
 | 
|  |   2976 |   Scan.optional
 | 
|  |   2977 |     (@{keyword "("} -- Parse.!!! (@{keyword "permissive"} -- @{keyword ")"}) >> K false) true --
 | 
|  |   2978 |   Parse.name -- Parse.term -- Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)
 | 
|  |   2979 |     -- Scan.option hints
 | 
|  |   2980 |   >> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map Parse.triple_swap eqs) src);
 | 
|  |   2981 | 
 | 
|  |   2982 | val _ =
 | 
|  |   2983 |   Outer_Syntax.command @{command_keyword recdef} "define general recursive functions (obsolete TFL)"
 | 
|  |   2984 |     (recdef_decl >> Toplevel.theory);
 | 
|  |   2985 | 
 | 
|  |   2986 | end;
 |