berghofe@5177: (* Title: HOL/Tools/datatype_abs_proofs.ML wenzelm@11539: Author: Stefan Berghofer, TU Muenchen berghofe@5177: berghofe@5177: Proofs and defintions independent of concrete representation berghofe@5177: of datatypes (i.e. requiring only abstract properties such as berghofe@5177: injectivity / distinctness of constructors and induction) berghofe@5177: berghofe@5177: - case distinction (exhaustion) theorems berghofe@5177: - characteristic equations for primrec combinators berghofe@5177: - characteristic equations for case combinators berghofe@5177: - equations for splitting "P (case ...)" expressions wenzelm@29264: - "nchotomy" and "case_cong" theorems for TFL berghofe@5177: *) berghofe@5177: berghofe@5177: signature DATATYPE_ABS_PROOFS = berghofe@5177: sig haftmann@31737: include DATATYPE_COMMON haftmann@31737: val prove_casedist_thms : config -> string list -> haftmann@31668: descr list -> (string * sort) list -> thm -> wenzelm@18728: attribute list -> theory -> thm list * theory haftmann@31737: val prove_primrec_thms : config -> string list -> haftmann@31668: descr list -> (string * sort) list -> haftmann@32915: (string -> thm list) -> thm list list -> thm list list * thm list list -> haftmann@32915: thm -> theory -> (string list * thm list) * theory haftmann@31737: val prove_case_thms : config -> string list -> haftmann@31668: descr list -> (string * sort) list -> haftmann@18314: string list -> thm list -> theory -> (thm list list * string list) * theory haftmann@31737: val prove_split_thms : config -> string list -> haftmann@31668: descr list -> (string * sort) list -> berghofe@5177: thm list list -> thm list list -> thm list -> thm list list -> theory -> haftmann@18314: (thm * thm) list * theory haftmann@31737: val prove_nchotomys : config -> string list -> descr list -> haftmann@18314: (string * sort) list -> thm list -> theory -> thm list * theory haftmann@31668: val prove_weak_case_congs : string list -> descr list -> haftmann@18314: (string * sort) list -> theory -> thm list * theory berghofe@13641: val prove_case_congs : string list -> haftmann@31668: descr list -> (string * sort) list -> haftmann@18314: thm list -> thm list list -> theory -> thm list * theory berghofe@5177: end; berghofe@5177: wenzelm@8436: structure DatatypeAbsProofs: DATATYPE_ABS_PROOFS = berghofe@5177: struct berghofe@5177: berghofe@5177: open DatatypeAux; berghofe@5177: berghofe@5177: (************************ case distinction theorems ***************************) berghofe@5177: haftmann@31737: fun prove_casedist_thms (config : config) new_type_names descr sorts induct case_names_exhausts thy = berghofe@5177: let haftmann@31668: val _ = message config "Proving case distinction theorems ..."; berghofe@5177: wenzelm@32952: val descr' = flat descr; berghofe@5177: val recTs = get_rec_types descr' sorts; haftmann@33955: val newTs = (uncurry take) (length (hd descr), recTs); berghofe@5177: berghofe@8477: val {maxidx, ...} = rep_thm induct; wenzelm@8305: val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct))); berghofe@5177: haftmann@33063: fun prove_casedist_thm (i, (T, t)) = berghofe@5177: let berghofe@5177: val dummyPs = map (fn (Var (_, Type (_, [T', T'']))) => berghofe@5177: Abs ("z", T', Const ("True", T''))) induct_Ps; berghofe@8477: val P = Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", maxidx+1), T), Bound 0) $ berghofe@5177: Var (("P", 0), HOLogic.boolT)) haftmann@33955: val insts = (uncurry take) (i, dummyPs) @ (P::((uncurry drop) (i + 1, dummyPs))); wenzelm@17985: val cert = cterm_of thy; berghofe@5177: val insts' = (map cert induct_Ps) ~~ (map cert insts); haftmann@32905: val induct' = refl RS ((nth haftmann@32905: (split_conj_thm (cterm_instantiate insts' induct)) i) RSN (2, rev_mp)) berghofe@5177: wenzelm@17985: in wenzelm@32970: Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t) wenzelm@26711: (fn {prems, ...} => EVERY wenzelm@17985: [rtac induct' 1, wenzelm@17985: REPEAT (rtac TrueI 1), wenzelm@17985: REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)), wenzelm@20046: REPEAT (rtac TrueI 1)]) berghofe@5177: end; berghofe@5177: haftmann@33063: val casedist_thms = map_index prove_casedist_thm haftmann@33063: (newTs ~~ DatatypeProp.make_casedists descr sorts) haftmann@18314: in haftmann@18314: thy haftmann@18314: |> store_thms_atts "exhaust" new_type_names (map single case_names_exhausts) casedist_thms haftmann@18314: end; berghofe@5177: berghofe@5177: berghofe@5177: (*************************** primrec combinators ******************************) berghofe@5177: haftmann@31737: fun prove_primrec_thms (config : config) new_type_names descr sorts haftmann@32915: injects_of constr_inject (dist_rewrites, other_dist_rewrites) induct thy = berghofe@5177: let haftmann@31668: val _ = message config "Constructing primrec combinators ..."; berghofe@5661: berghofe@5661: val big_name = space_implode "_" new_type_names; haftmann@32124: val thy0 = Sign.add_path big_name thy; berghofe@5177: wenzelm@32952: val descr' = flat descr; berghofe@5177: val recTs = get_rec_types descr' sorts; wenzelm@30190: val used = List.foldr OldTerm.add_typ_tfree_names [] recTs; haftmann@33955: val newTs = (uncurry take) (length (hd descr), recTs); berghofe@5177: wenzelm@8305: val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct))); berghofe@5177: berghofe@5661: val big_rec_name' = big_name ^ "_rec_set"; berghofe@21021: val rec_set_names' = berghofe@21021: if length descr' = 1 then [big_rec_name'] else berghofe@21021: map ((curry (op ^) (big_rec_name' ^ "_")) o string_of_int) berghofe@21021: (1 upto (length descr')); haftmann@28965: val rec_set_names = map (Sign.full_bname thy0) rec_set_names'; berghofe@5177: berghofe@15459: val (rec_result_Ts, reccomb_fn_Ts) = DatatypeProp.make_primrec_Ts descr sorts used; berghofe@5177: berghofe@21021: val rec_set_Ts = map (fn (T1, T2) => berghofe@21021: reccomb_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts); berghofe@5177: berghofe@5177: val rec_fns = map (uncurry (mk_Free "f")) berghofe@5177: (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts))); berghofe@21021: val rec_sets' = map (fn c => list_comb (Free c, rec_fns)) berghofe@21021: (rec_set_names' ~~ rec_set_Ts); berghofe@5177: val rec_sets = map (fn c => list_comb (Const c, rec_fns)) berghofe@5177: (rec_set_names ~~ rec_set_Ts); berghofe@5177: berghofe@5177: (* introduction rules for graph of primrec function *) berghofe@5177: haftmann@32906: fun make_rec_intr T rec_set (cname, cargs) (rec_intr_ts, l) = berghofe@5177: let wenzelm@33338: fun mk_prem (dt, U) (j, k, prems, t1s, t2s) = berghofe@7015: let val free1 = mk_Free "x" U j berghofe@13641: in (case (strip_dtyp dt, strip_type U) of berghofe@13641: ((_, DtRec m), (Us, _)) => berghofe@13641: let haftmann@32905: val free2 = mk_Free "y" (Us ---> nth rec_result_Ts m) k; berghofe@13641: val i = length Us berghofe@13641: in (j + 1, k + 1, HOLogic.mk_Trueprop (HOLogic.list_all haftmann@32905: (map (pair "x") Us, nth rec_sets' m $ berghofe@21021: app_bnds free1 i $ app_bnds free2 i)) :: prems, berghofe@5177: free1::t1s, free2::t2s) berghofe@5177: end berghofe@5177: | _ => (j + 1, k, prems, free1::t1s, t2s)) berghofe@5177: end; berghofe@5177: berghofe@5177: val Ts = map (typ_of_dtyp descr' sorts) cargs; wenzelm@33338: val (_, _, prems, t1s, t2s) = fold_rev mk_prem (cargs ~~ Ts) (1, 1, [], [], []) berghofe@5177: berghofe@21021: in (rec_intr_ts @ [Logic.list_implies (prems, HOLogic.mk_Trueprop berghofe@21021: (rec_set $ list_comb (Const (cname, Ts ---> T), t1s) $ haftmann@32905: list_comb (nth rec_fns l, t1s @ t2s)))], l + 1) berghofe@5177: end; berghofe@5177: haftmann@32906: val (rec_intr_ts, _) = fold (fn ((d, T), set_name) => haftmann@32906: fold (make_rec_intr T set_name) (#3 (snd d))) haftmann@32906: (descr' ~~ recTs ~~ rec_sets') ([], 0); berghofe@5177: wenzelm@21365: val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) = wenzelm@33278: thy0 wenzelm@33278: |> Sign.map_naming Name_Space.conceal wenzelm@33726: |> Inductive.add_inductive_global wenzelm@33669: {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name', wenzelm@33669: coind = false, no_elim = false, no_ind = true, skip_mono = true, fork_mono = false} haftmann@28965: (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts)) wenzelm@26128: (map dest_Free rec_fns) wenzelm@33278: (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] wenzelm@33278: ||> Sign.restore_naming thy0 wenzelm@33278: ||> Theory.checkpoint; berghofe@5177: berghofe@5177: (* prove uniqueness and termination of primrec combinators *) berghofe@5177: haftmann@31668: val _ = message config "Proving termination and uniqueness of primrec functions ..."; berghofe@5177: haftmann@32906: fun mk_unique_tac ((((i, (tname, _, constrs)), elim), T), T') (tac, intrs) = berghofe@5177: let berghofe@21021: val distinct_tac = berghofe@5177: (if i < length newTs then haftmann@32905: full_simp_tac (HOL_ss addsimps (nth dist_rewrites i)) 1 haftmann@32915: else full_simp_tac (HOL_ss addsimps (flat other_dist_rewrites)) 1); berghofe@5177: berghofe@5177: val inject = map (fn r => r RS iffD1) haftmann@32905: (if i < length newTs then nth constr_inject i haftmann@32729: else injects_of tname); berghofe@5177: haftmann@32906: fun mk_unique_constr_tac n (cname, cargs) (tac, intr::intrs, j) = berghofe@5177: let haftmann@32729: val k = length (filter is_rec_type cargs) berghofe@5177: berghofe@5177: in (EVERY [DETERM tac, berghofe@5177: REPEAT (etac ex1E 1), rtac ex1I 1, berghofe@5177: DEPTH_SOLVE_1 (ares_tac [intr] 1), berghofe@13641: REPEAT_DETERM_N k (etac thin_rl 1 THEN rotate_tac 1 1), berghofe@5177: etac elim 1, berghofe@5177: REPEAT_DETERM_N j distinct_tac, berghofe@21021: TRY (dresolve_tac inject 1), berghofe@5177: REPEAT (etac conjE 1), hyp_subst_tac 1, berghofe@13641: REPEAT (EVERY [etac allE 1, dtac mp 1, atac 1]), berghofe@5177: TRY (hyp_subst_tac 1), berghofe@5177: rtac refl 1, berghofe@5177: REPEAT_DETERM_N (n - j - 1) distinct_tac], berghofe@5177: intrs, j + 1) berghofe@5177: end; berghofe@5177: haftmann@32906: val (tac', intrs', _) = fold (mk_unique_constr_tac (length constrs)) haftmann@32906: constrs (tac, intrs, 0); berghofe@5177: berghofe@5177: in (tac', intrs') end; berghofe@5177: berghofe@5177: val rec_unique_thms = berghofe@5177: let berghofe@5177: val rec_unique_ts = map (fn (((set_t, T1), T2), i) => berghofe@5177: Const ("Ex1", (T2 --> HOLogic.boolT) --> HOLogic.boolT) $ berghofe@21021: absfree ("y", T2, set_t $ mk_Free "x" T1 i $ Free ("y", T2))) berghofe@21021: (rec_sets ~~ recTs ~~ rec_result_Ts ~~ (1 upto length recTs)); wenzelm@17985: val cert = cterm_of thy1 berghofe@5177: val insts = map (fn ((i, T), t) => absfree ("x" ^ (string_of_int i), T, t)) berghofe@5177: ((1 upto length recTs) ~~ recTs ~~ rec_unique_ts); berghofe@5177: val induct' = cterm_instantiate ((map cert induct_Ps) ~~ berghofe@5177: (map cert insts)) induct; haftmann@32906: val (tac, _) = fold mk_unique_tac (descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts) haftmann@32906: (((rtac induct' THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1 haftmann@32906: THEN rewrite_goals_tac [mk_meta_eq choice_eq], rec_intrs)); berghofe@5177: wenzelm@32970: in split_conj_thm (Skip_Proof.prove_global thy1 [] [] wenzelm@20046: (HOLogic.mk_Trueprop (mk_conj rec_unique_ts)) (K tac)) berghofe@5177: end; berghofe@5177: wenzelm@11435: val rec_total_thms = map (fn r => r RS theI') rec_unique_thms; berghofe@5177: berghofe@5177: (* define primrec combinators *) berghofe@5177: berghofe@5177: val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec"; haftmann@28965: val reccomb_names = map (Sign.full_bname thy1) berghofe@5177: (if length descr' = 1 then [big_reccomb_name] else berghofe@5177: (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int) berghofe@5177: (1 upto (length descr')))); berghofe@5177: val reccombs = map (fn ((name, T), T') => list_comb berghofe@5177: (Const (name, reccomb_fn_Ts @ [T] ---> T'), rec_fns)) berghofe@5177: (reccomb_names ~~ recTs ~~ rec_result_Ts); berghofe@5177: haftmann@18358: val (reccomb_defs, thy2) = haftmann@18358: thy1 wenzelm@24712: |> Sign.add_consts_i (map (fn ((name, T), T') => wenzelm@30364: (Binding.name (Long_Name.base_name name), reccomb_fn_Ts @ [T] ---> T', NoSyn)) haftmann@18358: (reccomb_names ~~ recTs ~~ rec_result_Ts)) haftmann@27691: |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') => wenzelm@30364: (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T, wenzelm@11435: Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T', berghofe@21021: set $ Free ("x", T) $ Free ("y", T')))))) haftmann@18358: (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts)) haftmann@32124: ||> Sign.parent_path wenzelm@28361: ||> Theory.checkpoint; berghofe@5177: berghofe@5177: berghofe@5177: (* prove characteristic equations for primrec combinators *) berghofe@5177: haftmann@31668: val _ = message config "Proving characteristic theorems for primrec combinators ..." berghofe@5177: wenzelm@32970: val rec_thms = map (fn t => Skip_Proof.prove_global thy2 [] [] t wenzelm@17985: (fn _ => EVERY wenzelm@17985: [rewrite_goals_tac reccomb_defs, wenzelm@17985: rtac the1_equality 1, berghofe@5177: resolve_tac rec_unique_thms 1, berghofe@5177: resolve_tac rec_intrs 1, wenzelm@20046: REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)])) berghofe@5177: (DatatypeProp.make_primrecs new_type_names descr sorts thy2) berghofe@5177: berghofe@5177: in haftmann@18314: thy2 wenzelm@24712: |> Sign.add_path (space_implode "_" new_type_names) blanchet@33056: |> PureThy.add_thmss [((Binding.name "recs", rec_thms), [Nitpick_Simps.add])] wenzelm@24712: ||> Sign.parent_path wenzelm@28361: ||> Theory.checkpoint haftmann@32906: |-> (fn thms => pair (reccomb_names, flat thms)) berghofe@5177: end; berghofe@5177: berghofe@8477: berghofe@5177: (***************************** case combinators *******************************) berghofe@5177: haftmann@31737: fun prove_case_thms (config : config) new_type_names descr sorts reccomb_names primrec_thms thy = berghofe@5177: let haftmann@31668: val _ = message config "Proving characteristic theorems for case combinators ..."; berghofe@5661: haftmann@32124: val thy1 = Sign.add_path (space_implode "_" new_type_names) thy; berghofe@5177: wenzelm@32952: val descr' = flat descr; berghofe@5177: val recTs = get_rec_types descr' sorts; wenzelm@30190: val used = List.foldr OldTerm.add_typ_tfree_names [] recTs; haftmann@33955: val newTs = (uncurry take) (length (hd descr), recTs); wenzelm@20071: val T' = TFree (Name.variant used "'t", HOLogic.typeS); berghofe@5177: berghofe@13641: fun mk_dummyT dt = binder_types (typ_of_dtyp descr' sorts dt) ---> T'; berghofe@7015: berghofe@5177: val case_dummy_fns = map (fn (_, (_, _, constrs)) => map (fn (_, cargs) => berghofe@5177: let berghofe@5177: val Ts = map (typ_of_dtyp descr' sorts) cargs; wenzelm@33317: val Ts' = map mk_dummyT (filter is_rec_type cargs) haftmann@28524: in Const (@{const_name undefined}, Ts @ Ts' ---> T') berghofe@5177: end) constrs) descr'; berghofe@5177: haftmann@28965: val case_names = map (fn s => Sign.full_bname thy1 (s ^ "_case")) new_type_names; berghofe@5177: berghofe@5177: (* define case combinators via primrec combinators *) berghofe@5177: haftmann@32906: val (case_defs, thy2) = fold (fn ((((i, (_, _, constrs)), T), name), recname) => fn (defs, thy) => berghofe@5177: let haftmann@32906: val (fns1, fns2) = split_list (map (fn ((_, cargs), j) => berghofe@5177: let berghofe@5177: val Ts = map (typ_of_dtyp descr' sorts) cargs; wenzelm@33317: val Ts' = Ts @ map mk_dummyT (filter is_rec_type cargs); berghofe@5177: val frees' = map (uncurry (mk_Free "x")) (Ts' ~~ (1 upto length Ts')); haftmann@33955: val frees = (uncurry take) (length cargs, frees'); berghofe@5177: val free = mk_Free "f" (Ts ---> T') j berghofe@5177: in berghofe@5177: (free, list_abs_free (map dest_Free frees', berghofe@5177: list_comb (free, frees))) berghofe@5177: end) (constrs ~~ (1 upto length constrs))); berghofe@5177: berghofe@5177: val caseT = (map (snd o dest_Free) fns1) @ [T] ---> T'; haftmann@33955: val fns = flat ((uncurry take) (i, case_dummy_fns)) @ haftmann@33955: fns2 @ flat ((uncurry drop) (i + 1, case_dummy_fns)); berghofe@5177: val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T'); wenzelm@30364: val decl = ((Binding.name (Long_Name.base_name name), caseT), NoSyn); wenzelm@30364: val def = (Binding.name (Long_Name.base_name name ^ "_def"), berghofe@5177: Logic.mk_equals (list_comb (Const (name, caseT), fns1), haftmann@33955: list_comb (reccomb, (flat ((uncurry take) (i, case_dummy_fns))) @ haftmann@33955: fns2 @ (flat ((uncurry drop) (i + 1, case_dummy_fns))) ))); haftmann@18358: val ([def_thm], thy') = haftmann@18358: thy wenzelm@33173: |> Sign.declare_const decl |> snd haftmann@27691: |> (PureThy.add_defs false o map Thm.no_attributes) [def]; berghofe@5177: wenzelm@8436: in (defs @ [def_thm], thy') haftmann@32906: end) (hd descr ~~ newTs ~~ case_names ~~ haftmann@33955: (uncurry take) (length newTs, reccomb_names)) ([], thy1) wenzelm@28361: ||> Theory.checkpoint; berghofe@5177: wenzelm@32970: val case_thms = map (map (fn t => Skip_Proof.prove_global thy2 [] [] t wenzelm@20046: (fn _ => EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1]))) berghofe@8477: (DatatypeProp.make_cases new_type_names descr sorts thy2) berghofe@8477: in haftmann@18314: thy2 wenzelm@33459: |> Context.theory_map ((fold o fold) Nitpick_Simps.add_thm case_thms) haftmann@32124: |> Sign.parent_path haftmann@18314: |> store_thmss "cases" new_type_names case_thms haftmann@18314: |-> (fn thmss => pair (thmss, case_names)) berghofe@8477: end; berghofe@5177: berghofe@5177: berghofe@5177: (******************************* case splitting *******************************) berghofe@5177: haftmann@31737: fun prove_split_thms (config : config) new_type_names descr sorts constr_inject dist_rewrites berghofe@5177: casedist_thms case_thms thy = berghofe@5177: let haftmann@31668: val _ = message config "Proving equations for case splitting ..."; berghofe@5177: haftmann@31668: val descr' = flat descr; berghofe@5177: val recTs = get_rec_types descr' sorts; haftmann@33955: val newTs = (uncurry take) (length (hd descr), recTs); berghofe@5177: berghofe@5177: fun prove_split_thms ((((((t1, t2), inject), dist_rewrites'), berghofe@5177: exhaustion), case_thms'), T) = berghofe@5177: let wenzelm@17985: val cert = cterm_of thy; berghofe@5177: val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (prems_of exhaustion))); berghofe@5177: val exhaustion' = cterm_instantiate berghofe@5177: [(cert lhs, cert (Free ("x", T)))] exhaustion; wenzelm@17985: val tacf = K (EVERY [rtac exhaustion' 1, ALLGOALS (asm_simp_tac wenzelm@17985: (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))]) berghofe@5177: in wenzelm@32970: (Skip_Proof.prove_global thy [] [] t1 tacf, wenzelm@32970: Skip_Proof.prove_global thy [] [] t2 tacf) berghofe@5177: end; berghofe@5177: berghofe@5177: val split_thm_pairs = map prove_split_thms berghofe@5177: ((DatatypeProp.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~ berghofe@5177: dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs); berghofe@5177: berghofe@5177: val (split_thms, split_asm_thms) = ListPair.unzip split_thm_pairs berghofe@5177: berghofe@5177: in haftmann@18314: thy haftmann@18314: |> store_thms "split" new_type_names split_thms haftmann@18314: ||>> store_thms "split_asm" new_type_names split_asm_thms haftmann@18314: |-> (fn (thms1, thms2) => pair (thms1 ~~ thms2)) berghofe@5177: end; berghofe@5177: nipkow@8601: fun prove_weak_case_congs new_type_names descr sorts thy = nipkow@8601: let nipkow@8601: fun prove_weak_case_cong t = wenzelm@32970: Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t) wenzelm@26711: (fn {prems, ...} => EVERY [rtac ((hd prems) RS arg_cong) 1]) nipkow@8601: nipkow@8601: val weak_case_congs = map prove_weak_case_cong (DatatypeProp.make_weak_case_congs nipkow@8601: new_type_names descr sorts thy) nipkow@8601: nipkow@8601: in thy |> store_thms "weak_case_cong" new_type_names weak_case_congs end; berghofe@8477: berghofe@5177: (************************* additional theorems for TFL ************************) berghofe@5177: haftmann@31737: fun prove_nchotomys (config : config) new_type_names descr sorts casedist_thms thy = berghofe@5177: let haftmann@31668: val _ = message config "Proving additional theorems for TFL ..."; berghofe@5177: berghofe@5177: fun prove_nchotomy (t, exhaustion) = berghofe@5177: let berghofe@5177: (* For goal i, select the correct disjunct to attack, then prove it *) berghofe@5177: fun tac i 0 = EVERY [TRY (rtac disjI1 i), berghofe@5177: hyp_subst_tac i, REPEAT (rtac exI i), rtac refl i] berghofe@5177: | tac i n = rtac disjI2 i THEN tac i (n - 1) berghofe@5177: in wenzelm@32970: Skip_Proof.prove_global thy [] [] t (fn _ => wenzelm@17985: EVERY [rtac allI 1, berghofe@5177: exh_tac (K exhaustion) 1, wenzelm@20046: ALLGOALS (fn i => tac i (i-1))]) berghofe@5177: end; berghofe@5177: berghofe@5177: val nchotomys = berghofe@5177: map prove_nchotomy (DatatypeProp.make_nchotomys descr sorts ~~ casedist_thms) berghofe@5177: wenzelm@8436: in thy |> store_thms "nchotomy" new_type_names nchotomys end; berghofe@5177: berghofe@5177: fun prove_case_congs new_type_names descr sorts nchotomys case_thms thy = berghofe@5177: let berghofe@5177: fun prove_case_cong ((t, nchotomy), case_rewrites) = berghofe@5177: let berghofe@5177: val (Const ("==>", _) $ tm $ _) = t; berghofe@5177: val (Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ Ma)) = tm; wenzelm@22578: val cert = cterm_of thy; berghofe@5177: val nchotomy' = nchotomy RS spec; wenzelm@29264: val [v] = Term.add_vars (concl_of nchotomy') []; wenzelm@29264: val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy' berghofe@5177: in wenzelm@32970: Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t) wenzelm@26711: (fn {prems, ...} => wenzelm@17985: let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites)) wenzelm@17985: in EVERY [simp_tac (HOL_ss addsimps [hd prems]) 1, wenzelm@17985: cut_facts_tac [nchotomy''] 1, wenzelm@17985: REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1), wenzelm@17985: REPEAT (etac exE 1) THEN simplify 1 (* Get last disjunct *)] wenzelm@20046: end) berghofe@5177: end; berghofe@5177: berghofe@5177: val case_congs = map prove_case_cong (DatatypeProp.make_case_congs berghofe@5177: new_type_names descr sorts thy ~~ nchotomys ~~ case_thms) berghofe@5177: wenzelm@8436: in thy |> store_thms "case_cong" new_type_names case_congs end; berghofe@5177: berghofe@5177: end;