berghofe@19494: (* Title: HOL/Nominal/nominal_package.ML berghofe@19494: ID: $Id$ berghofe@19494: Author: Stefan Berghofer and Christian Urban, TU Muenchen berghofe@19494: berghofe@19494: Nominal datatype package for Isabelle/HOL. berghofe@19494: *) berghofe@17870: berghofe@17870: signature NOMINAL_PACKAGE = berghofe@17870: sig berghofe@17870: val add_nominal_datatype : bool -> string list -> (string list * bstring * mixfix * berghofe@18068: (bstring * string list * mixfix) list) list -> theory -> theory berghofe@22433: type descr berghofe@21540: type nominal_datatype_info berghofe@21540: val get_nominal_datatypes : theory -> nominal_datatype_info Symtab.table berghofe@21540: val get_nominal_datatype : theory -> string -> nominal_datatype_info option berghofe@22311: val mk_perm: typ list -> term -> term -> term berghofe@22529: val perm_of_pair: term * term -> term berghofe@22529: val mk_not_sym: thm list -> thm list berghofe@22529: val perm_simproc: simproc berghofe@17870: end berghofe@17870: berghofe@18068: structure NominalPackage : NOMINAL_PACKAGE = berghofe@17870: struct berghofe@17870: berghofe@22274: val finite_emptyI = thm "finite.emptyI"; wenzelm@21669: val finite_Diff = thm "finite_Diff"; wenzelm@21669: val finite_Un = thm "finite_Un"; wenzelm@21669: val Un_iff = thm "Un_iff"; wenzelm@21669: val In0_eq = thm "In0_eq"; wenzelm@21669: val In1_eq = thm "In1_eq"; wenzelm@21669: val In0_not_In1 = thm "In0_not_In1"; wenzelm@21669: val In1_not_In0 = thm "In1_not_In0"; wenzelm@21669: val Un_assoc = thm "Un_assoc"; wenzelm@21669: val Collect_disj_eq = thm "Collect_disj_eq"; wenzelm@21669: val empty_def = thm "empty_def"; berghofe@24459: val empty_iff = thm "empty_iff"; wenzelm@21669: berghofe@17870: open DatatypeAux; berghofe@18068: open NominalAtoms; berghofe@17870: berghofe@18016: (** FIXME: DatatypePackage should export this function **) berghofe@18016: berghofe@18016: local berghofe@18016: berghofe@18016: fun dt_recs (DtTFree _) = [] berghofe@18016: | dt_recs (DtType (_, dts)) = List.concat (map dt_recs dts) berghofe@18016: | dt_recs (DtRec i) = [i]; berghofe@18016: berghofe@18016: fun dt_cases (descr: descr) (_, args, constrs) = berghofe@18016: let berghofe@18016: fun the_bname i = Sign.base_name (#1 (valOf (AList.lookup (op =) descr i))); berghofe@19133: val bnames = map the_bname (distinct op = (List.concat (map dt_recs args))); berghofe@18016: in map (fn (c, _) => space_implode "_" (Sign.base_name c :: bnames)) constrs end; berghofe@18016: berghofe@18016: berghofe@18016: fun induct_cases descr = berghofe@18016: DatatypeProp.indexify_names (List.concat (map (dt_cases descr) (map #2 descr))); berghofe@18016: berghofe@18016: fun exhaust_cases descr i = dt_cases descr (valOf (AList.lookup (op =) descr i)); berghofe@18016: berghofe@18016: in berghofe@18016: berghofe@18016: fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr); berghofe@18016: berghofe@18016: fun mk_case_names_exhausts descr new = berghofe@18016: map (RuleCases.case_names o exhaust_cases descr o #1) berghofe@18016: (List.filter (fn ((_, (name, _, _))) => name mem_string new) descr); berghofe@18016: berghofe@18016: end; berghofe@18016: wenzelm@22846: (* theory data *) berghofe@21540: berghofe@22433: type descr = (int * (string * dtyp list * (string * (dtyp list * dtyp) list) list)) list; berghofe@22433: berghofe@21540: type nominal_datatype_info = berghofe@21540: {index : int, berghofe@21540: descr : descr, berghofe@21540: sorts : (string * sort) list, berghofe@21540: rec_names : string list, berghofe@21540: rec_rewrites : thm list, berghofe@21540: induction : thm, berghofe@21540: distinct : thm list, berghofe@21540: inject : thm list}; berghofe@21540: berghofe@21540: structure NominalDatatypesData = TheoryDataFun wenzelm@22846: ( berghofe@21540: type T = nominal_datatype_info Symtab.table; berghofe@21540: val empty = Symtab.empty; berghofe@21540: val copy = I; berghofe@21540: val extend = I; berghofe@21540: fun merge _ tabs : T = Symtab.merge (K true) tabs; wenzelm@22846: ); berghofe@21540: berghofe@21540: val get_nominal_datatypes = NominalDatatypesData.get; berghofe@21540: val put_nominal_datatypes = NominalDatatypesData.put; berghofe@21540: val map_nominal_datatypes = NominalDatatypesData.map; berghofe@21540: val get_nominal_datatype = Symtab.lookup o get_nominal_datatypes; wenzelm@22846: berghofe@21540: berghofe@21540: (**** make datatype info ****) berghofe@21540: berghofe@21540: fun make_dt_info descr sorts induct reccomb_names rec_thms berghofe@21540: (((i, (_, (tname, _, _))), distinct), inject) = berghofe@21540: (tname, berghofe@21540: {index = i, berghofe@21540: descr = descr, berghofe@21540: sorts = sorts, berghofe@21540: rec_names = reccomb_names, berghofe@21540: rec_rewrites = rec_thms, berghofe@21540: induction = induct, berghofe@21540: distinct = distinct, berghofe@21540: inject = inject}); berghofe@21540: berghofe@18016: (*******************************) berghofe@18016: berghofe@17870: val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma); berghofe@17870: berghofe@17870: fun read_typ sign ((Ts, sorts), str) = berghofe@17870: let wenzelm@22675: val T = Type.no_tvars (Sign.read_def_typ (sign, (AList.lookup op =) berghofe@17870: (map (apfst (rpair ~1)) sorts)) str) handle TYPE (msg, _, _) => error msg berghofe@17870: in (Ts @ [T], add_typ_tfrees (T, sorts)) end; berghofe@17870: berghofe@17870: (** taken from HOL/Tools/datatype_aux.ML **) berghofe@17870: berghofe@17870: fun indtac indrule indnames i st = berghofe@17870: let berghofe@17870: val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule)); berghofe@17870: val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop berghofe@17870: (Logic.strip_imp_concl (List.nth (prems_of st, i - 1)))); berghofe@17870: val getP = if can HOLogic.dest_imp (hd ts) then berghofe@17870: (apfst SOME) o HOLogic.dest_imp else pair NONE; berghofe@17870: fun abstr (t1, t2) = (case t1 of berghofe@17870: NONE => (case filter (fn Free (s, _) => s mem indnames | _ => false) berghofe@17870: (term_frees t2) of berghofe@17870: [Free (s, T)] => absfree (s, T, t2) berghofe@17870: | _ => sys_error "indtac") berghofe@21021: | SOME (_ $ t') => Abs ("x", fastype_of t', abstract_over (t', t2))) wenzelm@22578: val cert = cterm_of (Thm.theory_of_thm st); berghofe@17870: val Ps = map (cert o head_of o snd o getP) ts; berghofe@17870: val indrule' = cterm_instantiate (Ps ~~ berghofe@17870: (map (cert o abstr o getP) ts')) indrule berghofe@17870: in berghofe@17870: rtac indrule' i st berghofe@17870: end; berghofe@17870: berghofe@18658: fun mk_subgoal i f st = berghofe@18658: let berghofe@18658: val cg = List.nth (cprems_of st, i - 1); berghofe@18658: val g = term_of cg; berghofe@18658: val revcut_rl' = Thm.lift_rule cg revcut_rl; berghofe@18658: val v = head_of (Logic.strip_assums_concl (hd (prems_of revcut_rl'))); berghofe@18658: val ps = Logic.strip_params g; wenzelm@22578: val cert = cterm_of (Thm.theory_of_thm st); berghofe@18658: in berghofe@18658: compose_tac (false, berghofe@18658: Thm.instantiate ([], [(cert v, cert (list_abs (ps, berghofe@18658: f (rev ps) (Logic.strip_assums_hyp g) (Logic.strip_assums_concl g))))]) berghofe@18658: revcut_rl', 2) i st berghofe@18658: end; berghofe@18658: berghofe@18658: (** simplification procedure for sorting permutations **) berghofe@18658: berghofe@18658: val dj_cp = thm "dj_cp"; berghofe@18658: berghofe@18658: fun dest_permT (Type ("fun", [Type ("List.list", [Type ("*", [T, _])]), berghofe@18658: Type ("fun", [_, U])])) = (T, U); berghofe@18658: berghofe@19494: fun permTs_of (Const ("Nominal.perm", T) $ t $ u) = fst (dest_permT T) :: permTs_of u berghofe@18658: | permTs_of _ = []; berghofe@18658: berghofe@19494: fun perm_simproc' thy ss (Const ("Nominal.perm", T) $ t $ (u as Const ("Nominal.perm", U) $ r $ s)) = berghofe@18658: let berghofe@18658: val (aT as Type (a, []), S) = dest_permT T; berghofe@18658: val (bT as Type (b, []), _) = dest_permT U berghofe@18658: in if aT mem permTs_of u andalso aT <> bT then berghofe@18658: let berghofe@18658: val a' = Sign.base_name a; berghofe@18658: val b' = Sign.base_name b; berghofe@18658: val cp = PureThy.get_thm thy (Name ("cp_" ^ a' ^ "_" ^ b' ^ "_inst")); berghofe@18658: val dj = PureThy.get_thm thy (Name ("dj_" ^ b' ^ "_" ^ a')); berghofe@18658: val dj_cp' = [cp, dj] MRS dj_cp; berghofe@18658: val cert = SOME o cterm_of thy berghofe@18658: in berghofe@18658: SOME (mk_meta_eq (Drule.instantiate' [SOME (ctyp_of thy S)] berghofe@18658: [cert t, cert r, cert s] dj_cp')) berghofe@18658: end berghofe@18658: else NONE berghofe@18658: end berghofe@18658: | perm_simproc' thy ss _ = NONE; berghofe@18658: berghofe@18658: val perm_simproc = berghofe@18658: Simplifier.simproc (the_context ()) "perm_simp" ["pi1 \\ (pi2 \\ x)"] perm_simproc'; berghofe@18658: berghofe@18658: val allE_Nil = read_instantiate_sg (the_context()) [("x", "[]")] allE; berghofe@18658: berghofe@18658: val meta_spec = thm "meta_spec"; berghofe@18658: wenzelm@18582: fun projections rule = wenzelm@19874: ProjectRule.projections (ProofContext.init (Thm.theory_of_thm rule)) rule wenzelm@18582: |> map (standard #> RuleCases.save rule); wenzelm@18582: berghofe@20267: val supp_prod = thm "supp_prod"; berghofe@20376: val fresh_prod = thm "fresh_prod"; berghofe@20376: val supports_fresh = thm "supports_fresh"; wenzelm@22812: val supports_def = thm "Nominal.supports_def"; berghofe@20376: val fresh_def = thm "fresh_def"; berghofe@20376: val supp_def = thm "supp_def"; berghofe@20376: val rev_simps = thms "rev.simps"; haftmann@23029: val app_simps = thms "append.simps"; berghofe@20267: berghofe@21021: val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq]; berghofe@21021: berghofe@22311: fun mk_perm Ts t u = berghofe@22311: let berghofe@22311: val T = fastype_of1 (Ts, t); berghofe@22311: val U = fastype_of1 (Ts, u) berghofe@22311: in Const ("Nominal.perm", T --> U --> U) $ t $ u end; berghofe@22311: berghofe@22529: fun perm_of_pair (x, y) = berghofe@22529: let berghofe@22529: val T = fastype_of x; berghofe@22529: val pT = mk_permT T berghofe@22529: in Const ("List.list.Cons", HOLogic.mk_prodT (T, T) --> pT --> pT) $ berghofe@22529: HOLogic.mk_prod (x, y) $ Const ("List.list.Nil", pT) berghofe@22529: end; berghofe@22529: berghofe@22529: fun mk_not_sym ths = maps (fn th => case prop_of th of berghofe@22529: _ $ (Const ("Not", _) $ _) => [th, th RS not_sym] berghofe@22529: | _ => [th]) ths; berghofe@22529: berghofe@17870: fun gen_add_nominal_datatype prep_typ err flat_names new_type_names dts thy = berghofe@17870: let berghofe@17870: (* this theory is used just for parsing *) wenzelm@21365: berghofe@17870: val tmp_thy = thy |> berghofe@17870: Theory.copy |> wenzelm@24712: Sign.add_types (map (fn (tvs, tname, mx, _) => berghofe@17870: (tname, length tvs, mx)) dts); berghofe@17870: berghofe@17870: val atoms = atoms_of thy; berghofe@17870: val classes = map (NameSpace.map_base (fn s => "pt_" ^ s)) atoms; berghofe@17870: val cp_classes = List.concat (map (fn atom1 => map (fn atom2 => berghofe@17870: Sign.intern_class thy ("cp_" ^ Sign.base_name atom1 ^ "_" ^ berghofe@17870: Sign.base_name atom2)) atoms) atoms); berghofe@17870: fun augment_sort S = S union classes; berghofe@17870: val augment_sort_typ = map_type_tfree (fn (s, S) => TFree (s, augment_sort S)); berghofe@17870: berghofe@17870: fun prep_constr ((constrs, sorts), (cname, cargs, mx)) = wenzelm@22578: let val (cargs', sorts') = Library.foldl (prep_typ tmp_thy) (([], sorts), cargs) berghofe@17870: in (constrs @ [(cname, cargs', mx)], sorts') end berghofe@17870: berghofe@17870: fun prep_dt_spec ((dts, sorts), (tvs, tname, mx, constrs)) = berghofe@17870: let val (constrs', sorts') = Library.foldl prep_constr (([], sorts), constrs) berghofe@17870: in (dts @ [(tvs, tname, mx, constrs')], sorts') end berghofe@17870: berghofe@17870: val (dts', sorts) = Library.foldl prep_dt_spec (([], []), dts); berghofe@17870: val sorts' = map (apsnd augment_sort) sorts; berghofe@17870: val tyvars = map #1 dts'; berghofe@17870: berghofe@17870: val types_syntax = map (fn (tvs, tname, mx, constrs) => (tname, mx)) dts'; berghofe@17870: val constr_syntax = map (fn (tvs, tname, mx, constrs) => berghofe@17870: map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts'; berghofe@17870: berghofe@17870: val ps = map (fn (_, n, _, _) => wenzelm@22578: (Sign.full_name tmp_thy n, Sign.full_name tmp_thy (n ^ "_Rep"))) dts; berghofe@17870: val rps = map Library.swap ps; berghofe@17870: wenzelm@21365: fun replace_types (Type ("Nominal.ABS", [T, U])) = berghofe@19494: Type ("fun", [T, Type ("Nominal.noption", [replace_types U])]) berghofe@17870: | replace_types (Type (s, Ts)) = berghofe@17870: Type (getOpt (AList.lookup op = ps s, s), map replace_types Ts) berghofe@17870: | replace_types T = T; berghofe@17870: berghofe@17870: val dts'' = map (fn (tvs, tname, mx, constrs) => (tvs, tname ^ "_Rep", NoSyn, berghofe@19833: map (fn (cname, cargs, mx) => (cname ^ "_Rep", berghofe@17870: map (augment_sort_typ o replace_types) cargs, NoSyn)) constrs)) dts'; berghofe@17870: berghofe@17870: val new_type_names' = map (fn n => n ^ "_Rep") new_type_names; wenzelm@22578: val full_new_type_names' = map (Sign.full_name thy) new_type_names'; berghofe@17870: urbanc@18045: val ({induction, ...},thy1) = berghofe@17870: DatatypePackage.add_datatype_i err flat_names new_type_names' dts'' thy; berghofe@17870: berghofe@17870: val SOME {descr, ...} = Symtab.lookup berghofe@17870: (DatatypePackage.get_datatypes thy1) (hd full_new_type_names'); berghofe@18107: fun nth_dtyp i = typ_of_dtyp descr sorts' (DtRec i); berghofe@18107: berghofe@17870: (**** define permutation functions ****) berghofe@17870: berghofe@17870: val permT = mk_permT (TFree ("'x", HOLogic.typeS)); berghofe@17870: val pi = Free ("pi", permT); berghofe@17870: val perm_types = map (fn (i, _) => berghofe@17870: let val T = nth_dtyp i berghofe@17870: in permT --> T --> T end) descr; berghofe@19494: val perm_names = replicate (length new_type_names) "Nominal.perm" @ wenzelm@22578: DatatypeProp.indexify_names (map (fn i => Sign.full_name thy1 berghofe@17870: ("perm_" ^ name_of_typ (nth_dtyp i))) berghofe@17870: (length new_type_names upto length descr - 1)); berghofe@17870: val perm_names_types = perm_names ~~ perm_types; berghofe@17870: berghofe@17870: val perm_eqs = List.concat (map (fn (i, (_, _, constrs)) => berghofe@17870: let val T = nth_dtyp i wenzelm@21365: in map (fn (cname, dts) => berghofe@17870: let berghofe@18107: val Ts = map (typ_of_dtyp descr sorts') dts; berghofe@17870: val names = DatatypeProp.make_tnames Ts; berghofe@17870: val args = map Free (names ~~ Ts); berghofe@17870: val c = Const (cname, Ts ---> T); berghofe@17870: fun perm_arg (dt, x) = berghofe@17870: let val T = type_of x berghofe@17870: in if is_rec_type dt then berghofe@17870: let val (Us, _) = strip_type T berghofe@17870: in list_abs (map (pair "x") Us, berghofe@17870: Const (List.nth (perm_names_types, body_index dt)) $ pi $ berghofe@17870: list_comb (x, map (fn (i, U) => berghofe@19494: Const ("Nominal.perm", permT --> U --> U) $ berghofe@17870: (Const ("List.rev", permT --> permT) $ pi) $ berghofe@17870: Bound i) ((length Us - 1 downto 0) ~~ Us))) berghofe@17870: end berghofe@19494: else Const ("Nominal.perm", permT --> T --> T) $ pi $ x wenzelm@21365: end; berghofe@17870: in berghofe@17870: (("", HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@17870: (Const (List.nth (perm_names_types, i)) $ berghofe@17870: Free ("pi", mk_permT (TFree ("'x", HOLogic.typeS))) $ berghofe@17870: list_comb (c, args), berghofe@17870: list_comb (c, map perm_arg (dts ~~ args))))), []) berghofe@17870: end) constrs berghofe@17870: end) descr); berghofe@17870: haftmann@20179: val (perm_simps, thy2) = thy1 |> wenzelm@24712: Sign.add_consts_i (map (fn (s, T) => (Sign.base_name s, T, NoSyn)) berghofe@17870: (List.drop (perm_names_types, length new_type_names))) |> wenzelm@19635: PrimrecPackage.add_primrec_unchecked_i "" perm_eqs; berghofe@17870: berghofe@17870: (**** prove that permutation functions introduced by unfolding are ****) berghofe@17870: (**** equivalent to already existing permutation functions ****) berghofe@17870: berghofe@17870: val _ = warning ("length descr: " ^ string_of_int (length descr)); berghofe@17870: val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names)); berghofe@17870: berghofe@17870: val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types); berghofe@17870: val perm_fun_def = PureThy.get_thm thy2 (Name "perm_fun_def"); berghofe@17870: berghofe@17870: val unfolded_perm_eq_thms = berghofe@17870: if length descr = length new_type_names then [] berghofe@17870: else map standard (List.drop (split_conj_thm wenzelm@20046: (Goal.prove_global thy2 [] [] berghofe@17870: (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj berghofe@17870: (map (fn (c as (s, T), x) => berghofe@17870: let val [T1, T2] = binder_types T berghofe@17870: in HOLogic.mk_eq (Const c $ pi $ Free (x, T2), berghofe@19494: Const ("Nominal.perm", T) $ pi $ Free (x, T2)) berghofe@17870: end) berghofe@18010: (perm_names_types ~~ perm_indnames)))) berghofe@18010: (fn _ => EVERY [indtac induction perm_indnames 1, berghofe@17870: ALLGOALS (asm_full_simp_tac berghofe@17870: (simpset_of thy2 addsimps [perm_fun_def]))])), berghofe@17870: length new_type_names)); berghofe@17870: berghofe@17870: (**** prove [] \ t = t ****) berghofe@17870: berghofe@17870: val _ = warning "perm_empty_thms"; berghofe@17870: berghofe@17870: val perm_empty_thms = List.concat (map (fn a => berghofe@17870: let val permT = mk_permT (Type (a, [])) berghofe@17870: in map standard (List.take (split_conj_thm wenzelm@20046: (Goal.prove_global thy2 [] [] berghofe@17870: (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj berghofe@17870: (map (fn ((s, T), x) => HOLogic.mk_eq berghofe@17870: (Const (s, permT --> T --> T) $ berghofe@17870: Const ("List.list.Nil", permT) $ Free (x, T), berghofe@17870: Free (x, T))) berghofe@17870: (perm_names ~~ berghofe@18010: map body_type perm_types ~~ perm_indnames)))) berghofe@18010: (fn _ => EVERY [indtac induction perm_indnames 1, berghofe@17870: ALLGOALS (asm_full_simp_tac (simpset_of thy2))])), berghofe@17870: length new_type_names)) berghofe@17870: end) berghofe@17870: atoms); berghofe@17870: berghofe@17870: (**** prove (pi1 @ pi2) \ t = pi1 \ (pi2 \ t) ****) berghofe@17870: berghofe@17870: val _ = warning "perm_append_thms"; berghofe@17870: berghofe@17870: (*FIXME: these should be looked up statically*) berghofe@17870: val at_pt_inst = PureThy.get_thm thy2 (Name "at_pt_inst"); berghofe@17870: val pt2 = PureThy.get_thm thy2 (Name "pt2"); berghofe@17870: berghofe@17870: val perm_append_thms = List.concat (map (fn a => berghofe@17870: let berghofe@17870: val permT = mk_permT (Type (a, [])); berghofe@17870: val pi1 = Free ("pi1", permT); berghofe@17870: val pi2 = Free ("pi2", permT); berghofe@17870: val pt_inst = PureThy.get_thm thy2 (Name ("pt_" ^ Sign.base_name a ^ "_inst")); berghofe@17870: val pt2' = pt_inst RS pt2; berghofe@17870: val pt2_ax = PureThy.get_thm thy2 berghofe@17870: (Name (NameSpace.map_base (fn s => "pt_" ^ s ^ "2") a)); berghofe@17870: in List.take (map standard (split_conj_thm wenzelm@20046: (Goal.prove_global thy2 [] [] berghofe@17870: (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj berghofe@17870: (map (fn ((s, T), x) => berghofe@17870: let val perm = Const (s, permT --> T --> T) berghofe@17870: in HOLogic.mk_eq haftmann@23029: (perm $ (Const ("List.append", permT --> permT --> permT) $ berghofe@17870: pi1 $ pi2) $ Free (x, T), berghofe@17870: perm $ pi1 $ (perm $ pi2 $ Free (x, T))) berghofe@17870: end) berghofe@17870: (perm_names ~~ berghofe@18010: map body_type perm_types ~~ perm_indnames)))) berghofe@18010: (fn _ => EVERY [indtac induction perm_indnames 1, berghofe@17870: ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt2', pt2_ax]))]))), berghofe@17870: length new_type_names) berghofe@17870: end) atoms); berghofe@17870: berghofe@17870: (**** prove pi1 ~ pi2 ==> pi1 \ t = pi2 \ t ****) berghofe@17870: berghofe@17870: val _ = warning "perm_eq_thms"; berghofe@17870: berghofe@17870: val pt3 = PureThy.get_thm thy2 (Name "pt3"); berghofe@17870: val pt3_rev = PureThy.get_thm thy2 (Name "pt3_rev"); berghofe@17870: berghofe@17870: val perm_eq_thms = List.concat (map (fn a => berghofe@17870: let berghofe@17870: val permT = mk_permT (Type (a, [])); berghofe@17870: val pi1 = Free ("pi1", permT); berghofe@17870: val pi2 = Free ("pi2", permT); berghofe@17870: (*FIXME: not robust - better access these theorems using NominalData?*) berghofe@17870: val at_inst = PureThy.get_thm thy2 (Name ("at_" ^ Sign.base_name a ^ "_inst")); berghofe@17870: val pt_inst = PureThy.get_thm thy2 (Name ("pt_" ^ Sign.base_name a ^ "_inst")); berghofe@17870: val pt3' = pt_inst RS pt3; berghofe@17870: val pt3_rev' = at_inst RS (pt_inst RS pt3_rev); berghofe@17870: val pt3_ax = PureThy.get_thm thy2 berghofe@17870: (Name (NameSpace.map_base (fn s => "pt_" ^ s ^ "3") a)); berghofe@17870: in List.take (map standard (split_conj_thm wenzelm@20046: (Goal.prove_global thy2 [] [] (Logic.mk_implies berghofe@19494: (HOLogic.mk_Trueprop (Const ("Nominal.prm_eq", berghofe@17870: permT --> permT --> HOLogic.boolT) $ pi1 $ pi2), berghofe@17870: HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj berghofe@17870: (map (fn ((s, T), x) => berghofe@17870: let val perm = Const (s, permT --> T --> T) berghofe@17870: in HOLogic.mk_eq berghofe@17870: (perm $ pi1 $ Free (x, T), berghofe@17870: perm $ pi2 $ Free (x, T)) berghofe@17870: end) berghofe@17870: (perm_names ~~ berghofe@18010: map body_type perm_types ~~ perm_indnames))))) berghofe@18010: (fn _ => EVERY [indtac induction perm_indnames 1, berghofe@17870: ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))), berghofe@17870: length new_type_names) berghofe@17870: end) atoms); berghofe@17870: berghofe@17870: (**** prove pi1 \ (pi2 \ t) = (pi1 \ pi2) \ (pi1 \ t) ****) berghofe@17870: berghofe@17870: val cp1 = PureThy.get_thm thy2 (Name "cp1"); berghofe@17870: val dj_cp = PureThy.get_thm thy2 (Name "dj_cp"); berghofe@17870: val pt_perm_compose = PureThy.get_thm thy2 (Name "pt_perm_compose"); berghofe@17870: val pt_perm_compose_rev = PureThy.get_thm thy2 (Name "pt_perm_compose_rev"); berghofe@17870: val dj_perm_perm_forget = PureThy.get_thm thy2 (Name "dj_perm_perm_forget"); berghofe@17870: berghofe@17870: fun composition_instance name1 name2 thy = berghofe@17870: let berghofe@17870: val name1' = Sign.base_name name1; berghofe@17870: val name2' = Sign.base_name name2; berghofe@17870: val cp_class = Sign.intern_class thy ("cp_" ^ name1' ^ "_" ^ name2'); berghofe@17870: val permT1 = mk_permT (Type (name1, [])); berghofe@17870: val permT2 = mk_permT (Type (name2, [])); berghofe@17870: val augment = map_type_tfree berghofe@17870: (fn (x, S) => TFree (x, cp_class :: S)); berghofe@17870: val Ts = map (augment o body_type) perm_types; berghofe@17870: val cp_inst = PureThy.get_thm thy berghofe@17870: (Name ("cp_" ^ name1' ^ "_" ^ name2' ^ "_inst")); berghofe@17870: val simps = simpset_of thy addsimps (perm_fun_def :: berghofe@17870: (if name1 <> name2 then berghofe@17870: let val dj = PureThy.get_thm thy (Name ("dj_" ^ name2' ^ "_" ^ name1')) berghofe@17870: in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end berghofe@17870: else berghofe@17870: let berghofe@17870: val at_inst = PureThy.get_thm thy (Name ("at_" ^ name1' ^ "_inst")); berghofe@17870: val pt_inst = PureThy.get_thm thy (Name ("pt_" ^ name1' ^ "_inst")) berghofe@17870: in berghofe@17870: [cp_inst RS cp1 RS sym, berghofe@17870: at_inst RS (pt_inst RS pt_perm_compose) RS sym, berghofe@17870: at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym] berghofe@17870: end)) wenzelm@20046: val thms = split_conj_thm (Goal.prove_global thy [] [] berghofe@17870: (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj berghofe@17870: (map (fn ((s, T), x) => berghofe@17870: let berghofe@17870: val pi1 = Free ("pi1", permT1); berghofe@17870: val pi2 = Free ("pi2", permT2); berghofe@17870: val perm1 = Const (s, permT1 --> T --> T); berghofe@17870: val perm2 = Const (s, permT2 --> T --> T); berghofe@19494: val perm3 = Const ("Nominal.perm", permT1 --> permT2 --> permT2) berghofe@17870: in HOLogic.mk_eq berghofe@17870: (perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)), berghofe@17870: perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T))) berghofe@17870: end) berghofe@18010: (perm_names ~~ Ts ~~ perm_indnames)))) berghofe@18010: (fn _ => EVERY [indtac induction perm_indnames 1, wenzelm@20046: ALLGOALS (asm_full_simp_tac simps)])) berghofe@17870: in berghofe@19275: foldl (fn ((s, tvs), thy) => AxClass.prove_arity berghofe@17870: (s, replicate (length tvs) (cp_class :: classes), [cp_class]) haftmann@24218: (Class.intro_classes_tac [] THEN ALLGOALS (resolve_tac thms)) thy) berghofe@17870: thy (full_new_type_names' ~~ tyvars) berghofe@17870: end; berghofe@17870: urbanc@18381: val (perm_thmss,thy3) = thy2 |> berghofe@17870: fold (fn name1 => fold (composition_instance name1) atoms) atoms |> berghofe@17870: curry (Library.foldr (fn ((i, (tyname, args, _)), thy) => berghofe@19275: AxClass.prove_arity (tyname, replicate (length args) classes, classes) haftmann@24218: (Class.intro_classes_tac [] THEN REPEAT (EVERY berghofe@17870: [resolve_tac perm_empty_thms 1, berghofe@17870: resolve_tac perm_append_thms 1, berghofe@17870: resolve_tac perm_eq_thms 1, assume_tac 1])) thy)) berghofe@17870: (List.take (descr, length new_type_names)) |> berghofe@17870: PureThy.add_thmss berghofe@17870: [((space_implode "_" new_type_names ^ "_unfolded_perm_eq", krauss@18759: unfolded_perm_eq_thms), [Simplifier.simp_add]), berghofe@17870: ((space_implode "_" new_type_names ^ "_perm_empty", krauss@18759: perm_empty_thms), [Simplifier.simp_add]), berghofe@17870: ((space_implode "_" new_type_names ^ "_perm_append", krauss@18759: perm_append_thms), [Simplifier.simp_add]), berghofe@17870: ((space_implode "_" new_type_names ^ "_perm_eq", krauss@18759: perm_eq_thms), [Simplifier.simp_add])]; wenzelm@21365: berghofe@17870: (**** Define representing sets ****) berghofe@17870: berghofe@17870: val _ = warning "representing sets"; berghofe@17870: berghofe@21021: val rep_set_names = DatatypeProp.indexify_names berghofe@21021: (map (fn (i, _) => name_of_typ (nth_dtyp i) ^ "_set") descr); berghofe@17870: val big_rep_name = berghofe@17870: space_implode "_" (DatatypeProp.indexify_names (List.mapPartial berghofe@19494: (fn (i, ("Nominal.noption", _, _)) => NONE berghofe@17870: | (i, _) => SOME (name_of_typ (nth_dtyp i))) descr)) ^ "_set"; berghofe@17870: val _ = warning ("big_rep_name: " ^ big_rep_name); berghofe@17870: berghofe@17870: fun strip_option (dtf as DtType ("fun", [dt, DtRec i])) = berghofe@17870: (case AList.lookup op = descr i of berghofe@19494: SOME ("Nominal.noption", _, [(_, [dt']), _]) => berghofe@17870: apfst (cons dt) (strip_option dt') berghofe@17870: | _ => ([], dtf)) berghofe@19494: | strip_option (DtType ("fun", [dt, DtType ("Nominal.noption", [dt'])])) = berghofe@18261: apfst (cons dt) (strip_option dt') berghofe@17870: | strip_option dt = ([], dt); berghofe@17870: berghofe@19133: val dt_atomTs = distinct op = (map (typ_of_dtyp descr sorts') berghofe@18280: (List.concat (map (fn (_, (_, _, cs)) => List.concat berghofe@18280: (map (List.concat o map (fst o strip_option) o snd) cs)) descr))); berghofe@18280: berghofe@17870: fun make_intr s T (cname, cargs) = berghofe@17870: let wenzelm@21365: fun mk_prem (dt, (j, j', prems, ts)) = berghofe@17870: let berghofe@17870: val (dts, dt') = strip_option dt; berghofe@17870: val (dts', dt'') = strip_dtyp dt'; berghofe@18107: val Ts = map (typ_of_dtyp descr sorts') dts; berghofe@18107: val Us = map (typ_of_dtyp descr sorts') dts'; berghofe@18107: val T = typ_of_dtyp descr sorts' dt''; berghofe@17870: val free = mk_Free "x" (Us ---> T) j; berghofe@17870: val free' = app_bnds free (length Us); berghofe@17870: fun mk_abs_fun (T, (i, t)) = berghofe@17870: let val U = fastype_of t berghofe@19494: in (i + 1, Const ("Nominal.abs_fun", [T, U, T] ---> berghofe@19494: Type ("Nominal.noption", [U])) $ mk_Free "y" T i $ t) berghofe@17870: end berghofe@17870: in (j + 1, j' + length Ts, berghofe@17870: case dt'' of berghofe@17870: DtRec k => list_all (map (pair "x") Us, berghofe@21021: HOLogic.mk_Trueprop (Free (List.nth (rep_set_names, k), berghofe@21021: T --> HOLogic.boolT) $ free')) :: prems berghofe@17870: | _ => prems, berghofe@17870: snd (foldr mk_abs_fun (j', free) Ts) :: ts) berghofe@17870: end; berghofe@17870: berghofe@17870: val (_, _, prems, ts) = foldr mk_prem (1, 1, [], []) cargs; berghofe@21021: val concl = HOLogic.mk_Trueprop (Free (s, T --> HOLogic.boolT) $ berghofe@21021: list_comb (Const (cname, map fastype_of ts ---> T), ts)) berghofe@17870: in Logic.list_implies (prems, concl) berghofe@17870: end; berghofe@17870: berghofe@21021: val (intr_ts, (rep_set_names', recTs')) = berghofe@21021: apfst List.concat (apsnd ListPair.unzip (ListPair.unzip (List.mapPartial berghofe@19494: (fn ((_, ("Nominal.noption", _, _)), _) => NONE berghofe@17870: | ((i, (_, _, constrs)), rep_set_name) => berghofe@17870: let val T = nth_dtyp i berghofe@17870: in SOME (map (make_intr rep_set_name T) constrs, berghofe@21021: (rep_set_name, T)) berghofe@17870: end) berghofe@21021: (descr ~~ rep_set_names)))); berghofe@21021: val rep_set_names'' = map (Sign.full_name thy3) rep_set_names'; berghofe@17870: wenzelm@21365: val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy4) = berghofe@17870: setmp InductivePackage.quiet_mode false wenzelm@24814: (InductivePackage.add_inductive_global {verbose = false, kind = Thm.internalK, wenzelm@24814: alt_name = big_rep_name, coind = false, no_elim = true, no_ind = false} berghofe@24746: (map (fn (s, T) => ((s, T --> HOLogic.boolT), NoSyn)) berghofe@21021: (rep_set_names' ~~ recTs')) berghofe@22607: [] (map (fn x => (("", []), x)) intr_ts) []) thy3; berghofe@17870: berghofe@17870: (**** Prove that representing set is closed under permutation ****) berghofe@17870: berghofe@17870: val _ = warning "proving closure under permutation..."; berghofe@17870: berghofe@17870: val perm_indnames' = List.mapPartial berghofe@19494: (fn (x, (_, ("Nominal.noption", _, _))) => NONE | (x, _) => SOME x) berghofe@17870: (perm_indnames ~~ descr); berghofe@17870: berghofe@17870: fun mk_perm_closed name = map (fn th => standard (th RS mp)) wenzelm@20046: (List.take (split_conj_thm (Goal.prove_global thy4 [] [] berghofe@17870: (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map berghofe@21021: (fn ((s, T), x) => berghofe@17870: let berghofe@21021: val T = map_type_tfree berghofe@21021: (fn (s, cs) => TFree (s, cs union cp_classes)) T; berghofe@21021: val S = Const (s, T --> HOLogic.boolT); berghofe@17870: val permT = mk_permT (Type (name, [])) berghofe@21021: in HOLogic.mk_imp (S $ Free (x, T), berghofe@21021: S $ (Const ("Nominal.perm", permT --> T --> T) $ berghofe@21021: Free ("pi", permT) $ Free (x, T))) berghofe@21021: end) (rep_set_names'' ~~ recTs' ~~ perm_indnames')))) berghofe@18010: (fn _ => EVERY (* CU: added perm_fun_def in the final tactic in order to deal with funs *) berghofe@17870: [indtac rep_induct [] 1, berghofe@17870: ALLGOALS (simp_tac (simpset_of thy4 addsimps berghofe@17870: (symmetric perm_fun_def :: PureThy.get_thms thy4 (Name ("abs_perm"))))), wenzelm@21365: ALLGOALS (resolve_tac rep_intrs berghofe@17870: THEN_ALL_NEW (asm_full_simp_tac (simpset_of thy4 addsimps [perm_fun_def])))])), berghofe@17870: length new_type_names)); berghofe@17870: berghofe@17870: (* FIXME: theorems are stored in database for testing only *) berghofe@17870: val perm_closed_thmss = map mk_perm_closed atoms; haftmann@20483: val (_, thy5) = PureThy.add_thmss [(("perm_closed", List.concat perm_closed_thmss), [])] thy4; berghofe@17870: berghofe@17870: (**** typedef ****) berghofe@17870: berghofe@17870: val _ = warning "defining type..."; berghofe@17870: berghofe@18366: val (typedefs, thy6) = haftmann@20483: thy5 berghofe@21021: |> fold_map (fn ((((name, mx), tvs), (cname, U)), name') => fn thy => berghofe@17870: setmp TypedefPackage.quiet_mode true berghofe@21021: (TypedefPackage.add_typedef_i false (SOME name') (name, tvs, mx) berghofe@21021: (Const ("Collect", (U --> HOLogic.boolT) --> HOLogic.mk_setT U) $ berghofe@21021: Const (cname, U --> HOLogic.boolT)) NONE berghofe@21021: (rtac exI 1 THEN rtac CollectI 1 THEN berghofe@17870: QUIET_BREADTH_FIRST (has_fewer_prems 1) haftmann@20483: (resolve_tac rep_intrs 1))) thy |> (fn ((_, r), thy) => berghofe@17870: let wenzelm@20071: val permT = mk_permT (TFree (Name.variant tvs "'a", HOLogic.typeS)); berghofe@17870: val pi = Free ("pi", permT); berghofe@17870: val tvs' = map (fn s => TFree (s, the (AList.lookup op = sorts' s))) tvs; berghofe@17870: val T = Type (Sign.intern_type thy name, tvs'); berghofe@18366: in apfst (pair r o hd) wenzelm@19635: (PureThy.add_defs_unchecked_i true [(("prm_" ^ name ^ "_def", Logic.mk_equals berghofe@19494: (Const ("Nominal.perm", permT --> T --> T) $ pi $ Free ("x", T), berghofe@17870: Const (Sign.intern_const thy ("Abs_" ^ name), U --> T) $ berghofe@19494: (Const ("Nominal.perm", permT --> U --> U) $ pi $ berghofe@17870: (Const (Sign.intern_const thy ("Rep_" ^ name), T --> U) $ berghofe@17870: Free ("x", T))))), [])] thy) berghofe@17870: end)) berghofe@18366: (types_syntax ~~ tyvars ~~ berghofe@21021: List.take (rep_set_names'' ~~ recTs', length new_type_names) ~~ berghofe@21021: new_type_names); berghofe@17870: berghofe@17870: val perm_defs = map snd typedefs; berghofe@21021: val Abs_inverse_thms = map (collect_simp o #Abs_inverse o fst) typedefs; berghofe@18016: val Rep_inverse_thms = map (#Rep_inverse o fst) typedefs; berghofe@21021: val Rep_thms = map (collect_simp o #Rep o fst) typedefs; berghofe@17870: berghofe@18016: val big_name = space_implode "_" new_type_names; berghofe@18016: berghofe@18016: berghofe@17870: (** prove that new types are in class pt_ **) berghofe@17870: berghofe@17870: val _ = warning "prove that new types are in class pt_ ..."; berghofe@17870: berghofe@17870: fun pt_instance ((class, atom), perm_closed_thms) = berghofe@21021: fold (fn ((((((Abs_inverse, Rep_inverse), Rep), berghofe@17870: perm_def), name), tvs), perm_closed) => fn thy => berghofe@19275: AxClass.prove_arity berghofe@17870: (Sign.intern_type thy name, berghofe@17870: replicate (length tvs) (classes @ cp_classes), [class]) haftmann@24218: (EVERY [Class.intro_classes_tac [], berghofe@17870: rewrite_goals_tac [perm_def], berghofe@17870: asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1, berghofe@17870: asm_full_simp_tac (simpset_of thy addsimps berghofe@17870: [Rep RS perm_closed RS Abs_inverse]) 1, berghofe@17870: asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy berghofe@17870: (Name ("pt_" ^ Sign.base_name atom ^ "3"))]) 1]) thy) berghofe@21021: (Abs_inverse_thms ~~ Rep_inverse_thms ~~ Rep_thms ~~ perm_defs ~~ berghofe@21021: new_type_names ~~ tyvars ~~ perm_closed_thms); berghofe@17870: berghofe@17870: berghofe@17870: (** prove that new types are in class cp__ **) berghofe@17870: berghofe@17870: val _ = warning "prove that new types are in class cp__ ..."; berghofe@17870: berghofe@17870: fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy = berghofe@17870: let berghofe@17870: val name = "cp_" ^ Sign.base_name atom1 ^ "_" ^ Sign.base_name atom2; berghofe@17870: val class = Sign.intern_class thy name; berghofe@17870: val cp1' = PureThy.get_thm thy (Name (name ^ "_inst")) RS cp1 berghofe@21021: in fold (fn ((((((Abs_inverse, Rep), berghofe@17870: perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy => berghofe@19275: AxClass.prove_arity berghofe@17870: (Sign.intern_type thy name, berghofe@17870: replicate (length tvs) (classes @ cp_classes), [class]) haftmann@24218: (EVERY [Class.intro_classes_tac [], berghofe@17870: rewrite_goals_tac [perm_def], berghofe@17870: asm_full_simp_tac (simpset_of thy addsimps berghofe@17870: ((Rep RS perm_closed1 RS Abs_inverse) :: berghofe@17870: (if atom1 = atom2 then [] berghofe@17870: else [Rep RS perm_closed2 RS Abs_inverse]))) 1, berghofe@18016: cong_tac 1, berghofe@17870: rtac refl 1, berghofe@17870: rtac cp1' 1]) thy) berghofe@21021: (Abs_inverse_thms ~~ Rep_thms ~~ perm_defs ~~ new_type_names ~~ berghofe@21021: tyvars ~~ perm_closed_thms1 ~~ perm_closed_thms2) thy berghofe@17870: end; berghofe@17870: berghofe@17870: val thy7 = fold (fn x => fn thy => thy |> berghofe@17870: pt_instance x |> berghofe@17870: fold (cp_instance (apfst snd x)) (atoms ~~ perm_closed_thmss)) berghofe@17870: (classes ~~ atoms ~~ perm_closed_thmss) thy6; berghofe@17870: berghofe@17870: (**** constructors ****) berghofe@17870: berghofe@17870: fun mk_abs_fun (x, t) = berghofe@17870: let berghofe@17870: val T = fastype_of x; berghofe@17870: val U = fastype_of t berghofe@17870: in berghofe@19494: Const ("Nominal.abs_fun", T --> U --> T --> berghofe@19494: Type ("Nominal.noption", [U])) $ x $ t berghofe@17870: end; berghofe@17870: berghofe@18016: val (ty_idxs, _) = foldl berghofe@19494: (fn ((i, ("Nominal.noption", _, _)), p) => p berghofe@18016: | ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr; berghofe@18016: berghofe@18016: fun reindex (DtType (s, dts)) = DtType (s, map reindex dts) berghofe@18016: | reindex (DtRec i) = DtRec (the (AList.lookup op = ty_idxs i)) berghofe@18016: | reindex dt = dt; berghofe@18016: berghofe@18016: fun strip_suffix i s = implode (List.take (explode s, size s - i)); berghofe@18016: berghofe@18016: (** strips the "_Rep" in type names *) wenzelm@21365: fun strip_nth_name i s = wenzelm@21858: let val xs = NameSpace.explode s; wenzelm@21858: in NameSpace.implode (Library.nth_map (length xs - i) (strip_suffix 4) xs) end; berghofe@18016: berghofe@18107: val (descr'', ndescr) = ListPair.unzip (List.mapPartial berghofe@19494: (fn (i, ("Nominal.noption", _, _)) => NONE berghofe@18107: | (i, (s, dts, constrs)) => berghofe@18107: let berghofe@18107: val SOME index = AList.lookup op = ty_idxs i; berghofe@18107: val (constrs1, constrs2) = ListPair.unzip berghofe@19833: (map (fn (cname, cargs) => apfst (pair (strip_nth_name 2 (strip_nth_name 1 cname))) berghofe@18107: (foldl_map (fn (dts, dt) => berghofe@18107: let val (dts', dt') = strip_option dt berghofe@18107: in (dts @ dts' @ [reindex dt'], (length dts, length dts')) end) berghofe@18107: ([], cargs))) constrs) berghofe@18107: in SOME ((index, (strip_nth_name 1 s, map reindex dts, constrs1)), berghofe@18107: (index, constrs2)) berghofe@18107: end) descr); urbanc@18045: berghofe@19489: val (descr1, descr2) = chop (length new_type_names) descr''; berghofe@18016: val descr' = [descr1, descr2]; berghofe@18016: berghofe@19710: fun partition_cargs idxs xs = map (fn (i, j) => berghofe@19710: (List.take (List.drop (xs, i), j), List.nth (xs, i + j))) idxs; berghofe@19710: berghofe@19833: val pdescr = map (fn ((i, (s, dts, constrs)), (_, idxss)) => (i, (s, dts, berghofe@19833: map (fn ((cname, cargs), idxs) => (cname, partition_cargs idxs cargs)) berghofe@19833: (constrs ~~ idxss)))) (descr'' ~~ ndescr); berghofe@19833: berghofe@19833: fun nth_dtyp' i = typ_of_dtyp descr'' sorts' (DtRec i); berghofe@17870: berghofe@17870: val rep_names = map (fn s => berghofe@17870: Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names; berghofe@17870: val abs_names = map (fn s => berghofe@17870: Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names; berghofe@17870: berghofe@18107: val recTs = get_rec_types descr'' sorts'; berghofe@18016: val newTs' = Library.take (length new_type_names, recTs'); berghofe@18016: val newTs = Library.take (length new_type_names, recTs); berghofe@17870: wenzelm@22578: val full_new_type_names = map (Sign.full_name thy) new_type_names; berghofe@17870: berghofe@19833: fun make_constr_def tname T T' ((thy, defs, eqns), berghofe@19833: (((cname_rep, _), (cname, cargs)), (cname', mx))) = berghofe@17870: let berghofe@19833: fun constr_arg ((dts, dt), (j, l_args, r_args)) = berghofe@17870: let berghofe@19833: val xs = map (fn (dt, i) => mk_Free "x" (typ_of_dtyp descr'' sorts' dt) i) berghofe@17870: (dts ~~ (j upto j + length dts - 1)) berghofe@19833: val x = mk_Free "x" (typ_of_dtyp descr'' sorts' dt) (j + length dts) berghofe@18261: in berghofe@18261: (j + length dts + 1, berghofe@18261: xs @ x :: l_args, berghofe@18261: foldr mk_abs_fun berghofe@19833: (case dt of berghofe@18261: DtRec k => if k < length new_type_names then berghofe@19833: Const (List.nth (rep_names, k), typ_of_dtyp descr'' sorts' dt --> berghofe@19833: typ_of_dtyp descr sorts' dt) $ x berghofe@18261: else error "nested recursion not (yet) supported" berghofe@18261: | _ => x) xs :: r_args) berghofe@17870: end berghofe@17870: berghofe@17870: val (_, l_args, r_args) = foldr constr_arg (1, [], []) cargs; wenzelm@22578: val abs_name = Sign.intern_const thy ("Abs_" ^ tname); wenzelm@22578: val rep_name = Sign.intern_const thy ("Rep_" ^ tname); berghofe@17870: val constrT = map fastype_of l_args ---> T; berghofe@19833: val lhs = list_comb (Const (cname, constrT), l_args); berghofe@19833: val rhs = list_comb (Const (cname_rep, map fastype_of r_args ---> T'), r_args); berghofe@17870: val def = Logic.mk_equals (lhs, Const (abs_name, T' --> T) $ rhs); berghofe@17870: val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@17870: (Const (rep_name, T --> T') $ lhs, rhs)); berghofe@17870: val def_name = (Sign.base_name cname) ^ "_def"; berghofe@18366: val ([def_thm], thy') = thy |> wenzelm@24712: Sign.add_consts_i [(cname', constrT, mx)] |> berghofe@17870: (PureThy.add_defs_i false o map Thm.no_attributes) [(def_name, def)] berghofe@17870: in (thy', defs @ [def_thm], eqns @ [eqn]) end; berghofe@17870: berghofe@19833: fun dt_constr_defs ((thy, defs, eqns, dist_lemmas), ((((((_, (_, _, constrs)), berghofe@19833: (_, (_, _, constrs'))), tname), T), T'), constr_syntax)) = berghofe@17870: let berghofe@17870: val rep_const = cterm_of thy berghofe@17870: (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T')); berghofe@17870: val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma); berghofe@17870: val (thy', defs', eqns') = Library.foldl (make_constr_def tname T T') wenzelm@24712: ((Sign.add_path tname thy, defs, []), constrs ~~ constrs' ~~ constr_syntax) berghofe@17870: in berghofe@17870: (parent_path flat_names thy', defs', eqns @ [eqns'], dist_lemmas @ [dist]) berghofe@17870: end; berghofe@17870: berghofe@17870: val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = Library.foldl dt_constr_defs berghofe@17870: ((thy7, [], [], []), List.take (descr, length new_type_names) ~~ berghofe@19833: List.take (pdescr, length new_type_names) ~~ berghofe@17870: new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax); berghofe@17870: berghofe@21021: val abs_inject_thms = map (collect_simp o #Abs_inject o fst) typedefs berghofe@21021: val rep_inject_thms = map (#Rep_inject o fst) typedefs berghofe@17870: berghofe@17870: (* prove theorem Rep_i (Constr_j ...) = Constr'_j ... *) wenzelm@21365: berghofe@17870: fun prove_constr_rep_thm eqn = berghofe@17870: let berghofe@17870: val inj_thms = map (fn r => r RS iffD1) abs_inject_thms; berghofe@21021: val rewrites = constr_defs @ map mk_meta_eq Rep_inverse_thms wenzelm@20046: in Goal.prove_global thy8 [] [] eqn (fn _ => EVERY berghofe@17870: [resolve_tac inj_thms 1, berghofe@17870: rewrite_goals_tac rewrites, berghofe@17870: rtac refl 3, berghofe@17870: resolve_tac rep_intrs 2, berghofe@21021: REPEAT (resolve_tac Rep_thms 1)]) berghofe@17870: end; berghofe@17870: berghofe@17870: val constr_rep_thmss = map (map prove_constr_rep_thm) constr_rep_eqns; berghofe@17870: berghofe@17870: (* prove theorem pi \ Rep_i x = Rep_i (pi \ x) *) berghofe@17870: berghofe@17870: fun prove_perm_rep_perm (atom, perm_closed_thms) = map (fn th => berghofe@17870: let berghofe@21021: val _ $ (_ $ (Rep $ x)) = Logic.unvarify (prop_of th); berghofe@17870: val Type ("fun", [T, U]) = fastype_of Rep; berghofe@17870: val permT = mk_permT (Type (atom, [])); berghofe@17870: val pi = Free ("pi", permT); berghofe@17870: in wenzelm@20046: Goal.prove_global thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@19494: (Const ("Nominal.perm", permT --> U --> U) $ pi $ (Rep $ x), berghofe@19494: Rep $ (Const ("Nominal.perm", permT --> T --> T) $ pi $ x)))) berghofe@18010: (fn _ => simp_tac (HOL_basic_ss addsimps (perm_defs @ Abs_inverse_thms @ wenzelm@20046: perm_closed_thms @ Rep_thms)) 1) berghofe@17870: end) Rep_thms; berghofe@17870: berghofe@17870: val perm_rep_perm_thms = List.concat (map prove_perm_rep_perm berghofe@17870: (atoms ~~ perm_closed_thmss)); berghofe@17870: berghofe@17870: (* prove distinctness theorems *) berghofe@17870: wenzelm@24110: val distinctness_limit = Config.get_thy thy8 DatatypeProp.distinctness_limit; wenzelm@24110: val thy8' = Config.put_thy DatatypeProp.distinctness_limit 1000 thy8; wenzelm@24098: val distinct_props = DatatypeProp.make_distincts new_type_names descr' sorts' thy8'; wenzelm@24110: val thy8 = Config.put_thy DatatypeProp.distinctness_limit distinctness_limit thy8'; berghofe@17870: berghofe@17870: val dist_rewrites = map (fn (rep_thms, dist_lemma) => berghofe@17870: dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0])) berghofe@17870: (constr_rep_thmss ~~ dist_lemmas); berghofe@17870: berghofe@17870: fun prove_distinct_thms (_, []) = [] berghofe@17870: | prove_distinct_thms (p as (rep_thms, dist_lemma), t::ts) = berghofe@17870: let wenzelm@20046: val dist_thm = Goal.prove_global thy8 [] [] t (fn _ => wenzelm@20046: simp_tac (simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1) berghofe@17870: in dist_thm::(standard (dist_thm RS not_sym)):: berghofe@17870: (prove_distinct_thms (p, ts)) berghofe@17870: end; berghofe@17870: berghofe@17870: val distinct_thms = map prove_distinct_thms berghofe@17870: (constr_rep_thmss ~~ dist_lemmas ~~ distinct_props); berghofe@17870: berghofe@17870: (** prove equations for permutation functions **) berghofe@17870: berghofe@17870: val abs_perm = PureThy.get_thms thy8 (Name "abs_perm"); (* FIXME *) berghofe@17870: berghofe@17870: val perm_simps' = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) => berghofe@19833: let val T = nth_dtyp' i berghofe@17870: in List.concat (map (fn (atom, perm_closed_thms) => wenzelm@21365: map (fn ((cname, dts), constr_rep_thm) => berghofe@17870: let berghofe@17870: val cname = Sign.intern_const thy8 berghofe@17870: (NameSpace.append tname (Sign.base_name cname)); berghofe@17870: val permT = mk_permT (Type (atom, [])); berghofe@17870: val pi = Free ("pi", permT); berghofe@17870: berghofe@17870: fun perm t = berghofe@17870: let val T = fastype_of t berghofe@19494: in Const ("Nominal.perm", permT --> T --> T) $ pi $ t end; berghofe@17870: berghofe@19833: fun constr_arg ((dts, dt), (j, l_args, r_args)) = berghofe@17870: let berghofe@19833: val Ts = map (typ_of_dtyp descr'' sorts') dts; berghofe@17870: val xs = map (fn (T, i) => mk_Free "x" T i) berghofe@17870: (Ts ~~ (j upto j + length dts - 1)) berghofe@19833: val x = mk_Free "x" (typ_of_dtyp descr'' sorts' dt) (j + length dts) berghofe@18261: in berghofe@18261: (j + length dts + 1, berghofe@18261: xs @ x :: l_args, berghofe@18261: map perm (xs @ [x]) @ r_args) berghofe@17870: end berghofe@17870: berghofe@17870: val (_, l_args, r_args) = foldr constr_arg (1, [], []) dts; berghofe@17870: val c = Const (cname, map fastype_of l_args ---> T) berghofe@17870: in wenzelm@20046: Goal.prove_global thy8 [] [] berghofe@17870: (HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@18010: (perm (list_comb (c, l_args)), list_comb (c, r_args)))) berghofe@18010: (fn _ => EVERY berghofe@17870: [simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1, berghofe@17870: simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @ berghofe@17870: constr_defs @ perm_closed_thms)) 1, berghofe@17870: TRY (simp_tac (HOL_basic_ss addsimps berghofe@17870: (symmetric perm_fun_def :: abs_perm)) 1), berghofe@17870: TRY (simp_tac (HOL_basic_ss addsimps berghofe@17870: (perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @ wenzelm@20046: perm_closed_thms)) 1)]) berghofe@17870: end) (constrs ~~ constr_rep_thms)) (atoms ~~ perm_closed_thmss)) berghofe@19833: end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss); berghofe@17870: berghofe@17870: (** prove injectivity of constructors **) berghofe@17870: berghofe@17870: val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms; berghofe@17870: val alpha = PureThy.get_thms thy8 (Name "alpha"); berghofe@17870: val abs_fresh = PureThy.get_thms thy8 (Name "abs_fresh"); berghofe@17870: berghofe@17870: val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) => berghofe@19833: let val T = nth_dtyp' i berghofe@17870: in List.mapPartial (fn ((cname, dts), constr_rep_thm) => berghofe@17870: if null dts then NONE else SOME berghofe@17870: let berghofe@17870: val cname = Sign.intern_const thy8 berghofe@17870: (NameSpace.append tname (Sign.base_name cname)); berghofe@17870: berghofe@19833: fun make_inj ((dts, dt), (j, args1, args2, eqs)) = berghofe@17870: let berghofe@19833: val Ts_idx = map (typ_of_dtyp descr'' sorts') dts ~~ (j upto j + length dts - 1); berghofe@17870: val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx; berghofe@17870: val ys = map (fn (T, i) => mk_Free "y" T i) Ts_idx; berghofe@19833: val x = mk_Free "x" (typ_of_dtyp descr'' sorts' dt) (j + length dts); berghofe@19833: val y = mk_Free "y" (typ_of_dtyp descr'' sorts' dt) (j + length dts) berghofe@18261: in berghofe@18261: (j + length dts + 1, berghofe@18261: xs @ (x :: args1), ys @ (y :: args2), berghofe@18261: HOLogic.mk_eq berghofe@18261: (foldr mk_abs_fun x xs, foldr mk_abs_fun y ys) :: eqs) berghofe@17870: end; berghofe@17870: berghofe@17870: val (_, args1, args2, eqs) = foldr make_inj (1, [], [], []) dts; berghofe@17870: val Ts = map fastype_of args1; berghofe@17870: val c = Const (cname, Ts ---> T) berghofe@17870: in wenzelm@20046: Goal.prove_global thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@17870: (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)), berghofe@18010: foldr1 HOLogic.mk_conj eqs))) berghofe@18010: (fn _ => EVERY berghofe@17870: [asm_full_simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: berghofe@17870: rep_inject_thms')) 1, berghofe@17870: TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def :: berghofe@17870: alpha @ abs_perm @ abs_fresh @ rep_inject_thms @ berghofe@17874: perm_rep_perm_thms)) 1), berghofe@17874: TRY (asm_full_simp_tac (HOL_basic_ss addsimps (perm_fun_def :: wenzelm@20046: expand_fun_eq :: rep_inject_thms @ perm_rep_perm_thms)) 1)]) berghofe@17870: end) (constrs ~~ constr_rep_thms) berghofe@19833: end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss); berghofe@17870: berghofe@17872: (** equations for support and freshness **) berghofe@17872: berghofe@17872: val (supp_thms, fresh_thms) = ListPair.unzip (map ListPair.unzip berghofe@17872: (map (fn ((((i, (_, _, constrs)), tname), inject_thms'), perm_thms') => berghofe@19833: let val T = nth_dtyp' i berghofe@17872: in List.concat (map (fn (cname, dts) => map (fn atom => berghofe@17872: let berghofe@17872: val cname = Sign.intern_const thy8 berghofe@17872: (NameSpace.append tname (Sign.base_name cname)); berghofe@17872: val atomT = Type (atom, []); berghofe@17872: berghofe@19833: fun process_constr ((dts, dt), (j, args1, args2)) = berghofe@17872: let berghofe@19833: val Ts_idx = map (typ_of_dtyp descr'' sorts') dts ~~ (j upto j + length dts - 1); berghofe@17872: val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx; berghofe@19833: val x = mk_Free "x" (typ_of_dtyp descr'' sorts' dt) (j + length dts) berghofe@18261: in berghofe@18261: (j + length dts + 1, berghofe@18261: xs @ (x :: args1), foldr mk_abs_fun x xs :: args2) berghofe@17872: end; berghofe@17872: berghofe@17872: val (_, args1, args2) = foldr process_constr (1, [], []) dts; berghofe@17872: val Ts = map fastype_of args1; berghofe@17872: val c = list_comb (Const (cname, Ts ---> T), args1); berghofe@17872: fun supp t = berghofe@19494: Const ("Nominal.supp", fastype_of t --> HOLogic.mk_setT atomT) $ t; berghofe@17872: fun fresh t = berghofe@19494: Const ("Nominal.fresh", atomT --> fastype_of t --> HOLogic.boolT) $ berghofe@17872: Free ("a", atomT) $ t; wenzelm@20046: val supp_thm = Goal.prove_global thy8 [] [] berghofe@17872: (HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@17872: (supp c, berghofe@17872: if null dts then Const ("{}", HOLogic.mk_setT atomT) berghofe@18010: else foldr1 (HOLogic.mk_binop "op Un") (map supp args2)))) berghofe@17872: (fn _ => berghofe@18010: simp_tac (HOL_basic_ss addsimps (supp_def :: berghofe@17872: Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un :: berghofe@22274: symmetric empty_def :: finite_emptyI :: simp_thms @ wenzelm@20046: abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1) berghofe@17872: in berghofe@17872: (supp_thm, wenzelm@20046: Goal.prove_global thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@17872: (fresh c, berghofe@17872: if null dts then HOLogic.true_const berghofe@18010: else foldr1 HOLogic.mk_conj (map fresh args2)))) berghofe@17872: (fn _ => berghofe@24459: simp_tac (HOL_ss addsimps [Un_iff, empty_iff, fresh_def, supp_thm]) 1)) berghofe@17872: end) atoms) constrs) berghofe@19833: end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps'))); berghofe@17872: berghofe@18107: (**** weak induction theorem ****) berghofe@18016: berghofe@18016: fun mk_indrule_lemma ((prems, concls), (((i, _), T), U)) = berghofe@18016: let berghofe@18016: val Rep_t = Const (List.nth (rep_names, i), T --> U) $ berghofe@18016: mk_Free "x" T i; berghofe@18016: berghofe@18016: val Abs_t = Const (List.nth (abs_names, i), U --> T) berghofe@18016: berghofe@21021: in (prems @ [HOLogic.imp $ berghofe@21021: (Const (List.nth (rep_set_names'', i), U --> HOLogic.boolT) $ Rep_t) $ berghofe@18016: (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))], berghofe@18016: concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i]) berghofe@18016: end; berghofe@18016: berghofe@18016: val (indrule_lemma_prems, indrule_lemma_concls) = berghofe@18107: Library.foldl mk_indrule_lemma (([], []), (descr'' ~~ recTs ~~ recTs')); berghofe@18016: wenzelm@20046: val indrule_lemma = Goal.prove_global thy8 [] [] berghofe@18016: (Logic.mk_implies berghofe@18016: (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems), berghofe@18016: HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY berghofe@18016: [REPEAT (etac conjE 1), berghofe@18016: REPEAT (EVERY berghofe@18016: [TRY (rtac conjI 1), full_simp_tac (HOL_basic_ss addsimps Rep_inverse_thms) 1, wenzelm@20046: etac mp 1, resolve_tac Rep_thms 1])]); berghofe@18016: berghofe@18016: val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma))); berghofe@18016: val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else berghofe@18016: map (Free o apfst fst o dest_Var) Ps; berghofe@18016: val indrule_lemma' = cterm_instantiate berghofe@18016: (map (cterm_of thy8) Ps ~~ map (cterm_of thy8) frees) indrule_lemma; berghofe@18016: berghofe@19833: val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms; berghofe@18016: berghofe@18016: val dt_induct_prop = DatatypeProp.make_ind descr' sorts'; wenzelm@20046: val dt_induct = Goal.prove_global thy8 [] berghofe@18016: (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop) berghofe@18016: (fn prems => EVERY berghofe@18016: [rtac indrule_lemma' 1, wenzelm@23590: (DatatypeAux.indtac rep_induct THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, berghofe@18016: EVERY (map (fn (prem, r) => (EVERY berghofe@18016: [REPEAT (eresolve_tac Abs_inverse_thms' 1), berghofe@18016: simp_tac (HOL_basic_ss addsimps [symmetric r]) 1, berghofe@18016: DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)])) wenzelm@20046: (prems ~~ constr_defs))]); berghofe@18016: berghofe@18107: val case_names_induct = mk_case_names_induct descr''; berghofe@18016: berghofe@18066: (**** prove that new datatypes have finite support ****) berghofe@18066: urbanc@18246: val _ = warning "proving finite support for the new datatype"; urbanc@18246: berghofe@18066: val indnames = DatatypeProp.make_tnames recTs; berghofe@18066: berghofe@18066: val abs_supp = PureThy.get_thms thy8 (Name "abs_supp"); urbanc@18067: val supp_atm = PureThy.get_thms thy8 (Name "supp_atm"); berghofe@18066: berghofe@18066: val finite_supp_thms = map (fn atom => berghofe@18066: let val atomT = Type (atom, []) berghofe@18066: in map standard (List.take wenzelm@20046: (split_conj_thm (Goal.prove_global thy8 [] [] (HOLogic.mk_Trueprop berghofe@22274: (foldr1 HOLogic.mk_conj (map (fn (s, T) => berghofe@22274: Const ("Finite_Set.finite", HOLogic.mk_setT atomT --> HOLogic.boolT) $ berghofe@22274: (Const ("Nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T))) berghofe@18066: (indnames ~~ recTs)))) berghofe@18066: (fn _ => indtac dt_induct indnames 1 THEN berghofe@18066: ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps urbanc@18067: (abs_supp @ supp_atm @ berghofe@18066: PureThy.get_thms thy8 (Name ("fs_" ^ Sign.base_name atom ^ "1")) @ berghofe@18066: List.concat supp_thms))))), berghofe@18066: length new_type_names)) berghofe@18066: end) atoms; berghofe@18066: krauss@18759: val simp_atts = replicate (length new_type_names) [Simplifier.simp_add]; berghofe@18658: urbanc@22245: (* Function to add both the simp and eqvt attributes *) urbanc@22245: (* These two attributes are duplicated on all the types in the mutual nominal datatypes *) urbanc@22245: urbanc@22245: val simp_eqvt_atts = replicate (length new_type_names) [Simplifier.simp_add, NominalThmDecls.eqvt_add]; urbanc@22245: berghofe@18658: val (_, thy9) = thy8 |> wenzelm@24712: Sign.add_path big_name |> berghofe@22543: PureThy.add_thms [(("weak_induct", dt_induct), [case_names_induct])] ||>> berghofe@22543: PureThy.add_thmss [(("weak_inducts", projections dt_induct), [case_names_induct])] ||> wenzelm@24712: Sign.parent_path ||>> berghofe@18658: DatatypeAux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>> berghofe@18658: DatatypeAux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>> urbanc@22231: DatatypeAux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>> berghofe@18658: DatatypeAux.store_thmss "inject" new_type_names inject_thms ||>> berghofe@18658: DatatypeAux.store_thmss "supp" new_type_names supp_thms ||>> berghofe@18658: DatatypeAux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||> berghofe@18658: fold (fn (atom, ths) => fn thy => berghofe@18658: let val class = Sign.intern_class thy ("fs_" ^ Sign.base_name atom) berghofe@19275: in fold (fn T => AxClass.prove_arity berghofe@18658: (fst (dest_Type T), berghofe@18658: replicate (length sorts) [class], [class]) haftmann@24218: (Class.intro_classes_tac [] THEN resolve_tac ths 1)) newTs thy berghofe@18658: end) (atoms ~~ finite_supp_thms); berghofe@18658: berghofe@18107: (**** strong induction theorem ****) berghofe@18107: berghofe@18107: val pnames = if length descr'' = 1 then ["P"] berghofe@18107: else map (fn i => "P" ^ string_of_int i) (1 upto length descr''); berghofe@18245: val ind_sort = if null dt_atomTs then HOLogic.typeS wenzelm@19649: else Sign.certify_sort thy9 (map (fn T => Sign.intern_class thy9 ("fs_" ^ berghofe@18658: Sign.base_name (fst (dest_Type T)))) dt_atomTs); berghofe@18107: val fsT = TFree ("'n", ind_sort); berghofe@18658: val fsT' = TFree ("'n", HOLogic.typeS); berghofe@18107: berghofe@18658: val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T))) berghofe@18658: (DatatypeProp.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs); berghofe@18658: berghofe@18658: fun make_pred fsT i T = berghofe@18302: Free (List.nth (pnames, i), fsT --> T --> HOLogic.boolT); berghofe@18107: berghofe@19851: fun mk_fresh1 xs [] = [] berghofe@19851: | mk_fresh1 xs ((y as (_, T)) :: ys) = map (fn x => HOLogic.mk_Trueprop berghofe@19851: (HOLogic.mk_not (HOLogic.mk_eq (Free y, Free x)))) berghofe@19851: (filter (fn (_, U) => T = U) (rev xs)) @ berghofe@19851: mk_fresh1 (y :: xs) ys; berghofe@19851: berghofe@19851: fun mk_fresh2 xss [] = [] berghofe@19851: | mk_fresh2 xss ((p as (ys, _)) :: yss) = List.concat (map (fn y as (_, T) => berghofe@19851: map (fn (_, x as (_, U)) => HOLogic.mk_Trueprop berghofe@19851: (Const ("Nominal.fresh", T --> U --> HOLogic.boolT) $ Free y $ Free x)) berghofe@19851: (rev xss @ yss)) ys) @ berghofe@19851: mk_fresh2 (p :: xss) yss; berghofe@19851: berghofe@18658: fun make_ind_prem fsT f k T ((cname, cargs), idxs) = berghofe@18107: let berghofe@18107: val recs = List.filter is_rec_type cargs; berghofe@18107: val Ts = map (typ_of_dtyp descr'' sorts') cargs; berghofe@18107: val recTs' = map (typ_of_dtyp descr'' sorts') recs; wenzelm@20071: val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts); berghofe@18107: val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs)); berghofe@18107: val frees = tnames ~~ Ts; berghofe@19710: val frees' = partition_cargs idxs frees; wenzelm@20071: val z = (Name.variant tnames "z", fsT); berghofe@18107: berghofe@18107: fun mk_prem ((dt, s), T) = berghofe@18107: let berghofe@18107: val (Us, U) = strip_type T; berghofe@18107: val l = length Us berghofe@18107: in list_all (z :: map (pair "x") Us, HOLogic.mk_Trueprop berghofe@18658: (make_pred fsT (body_index dt) U $ Bound l $ app_bnds (Free (s, T)) l)) berghofe@18107: end; berghofe@18107: berghofe@18107: val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs'); berghofe@18107: val prems' = map (fn p as (_, T) => HOLogic.mk_Trueprop berghofe@19710: (f T (Free p) (Free z))) (List.concat (map fst frees')) @ berghofe@19710: mk_fresh1 [] (List.concat (map fst frees')) @ berghofe@19710: mk_fresh2 [] frees' berghofe@18107: berghofe@18302: in list_all_free (frees @ [z], Logic.list_implies (prems' @ prems, berghofe@18658: HOLogic.mk_Trueprop (make_pred fsT k T $ Free z $ berghofe@18302: list_comb (Const (cname, Ts ---> T), map Free frees)))) berghofe@18107: end; berghofe@18107: berghofe@18107: val ind_prems = List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) => berghofe@18658: map (make_ind_prem fsT (fn T => fn t => fn u => berghofe@19494: Const ("Nominal.fresh", T --> fsT --> HOLogic.boolT) $ t $ u) i T) berghofe@18658: (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs)); berghofe@18107: val tnames = DatatypeProp.make_tnames recTs; wenzelm@20071: val zs = Name.variant_list tnames (replicate (length descr'') "z"); berghofe@18107: val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") berghofe@18658: (map (fn ((((i, _), T), tname), z) => berghofe@18658: make_pred fsT i T $ Free (z, fsT) $ Free (tname, T)) berghofe@18658: (descr'' ~~ recTs ~~ tnames ~~ zs))); berghofe@18107: val induct = Logic.list_implies (ind_prems, ind_concl); berghofe@18107: berghofe@18658: val ind_prems' = berghofe@18658: map (fn (_, f as Free (_, T)) => list_all_free ([("x", fsT')], berghofe@22274: HOLogic.mk_Trueprop (Const ("Finite_Set.finite", body_type T --> berghofe@22274: HOLogic.boolT) $ (f $ Free ("x", fsT'))))) fresh_fs @ berghofe@18658: List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) => berghofe@18658: map (make_ind_prem fsT' (fn T => fn t => fn u => HOLogic.Not $ berghofe@18658: HOLogic.mk_mem (t, the (AList.lookup op = fresh_fs T) $ u)) i T) berghofe@18658: (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs)); berghofe@18658: val ind_concl' = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") berghofe@18658: (map (fn ((((i, _), T), tname), z) => berghofe@18658: make_pred fsT' i T $ Free (z, fsT') $ Free (tname, T)) berghofe@18658: (descr'' ~~ recTs ~~ tnames ~~ zs))); berghofe@18658: val induct' = Logic.list_implies (ind_prems', ind_concl'); berghofe@18658: berghofe@18658: val aux_ind_vars = berghofe@18658: (DatatypeProp.indexify_names (replicate (length dt_atomTs) "pi") ~~ berghofe@18658: map mk_permT dt_atomTs) @ [("z", fsT')]; berghofe@18658: val aux_ind_Ts = rev (map snd aux_ind_vars); berghofe@18658: val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") berghofe@18658: (map (fn (((i, _), T), tname) => berghofe@18658: HOLogic.list_all (aux_ind_vars, make_pred fsT' i T $ Bound 0 $ berghofe@22311: fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1)) berghofe@22311: (Free (tname, T)))) berghofe@18658: (descr'' ~~ recTs ~~ tnames))); berghofe@18658: berghofe@18658: fun mk_ind_perm i k p l vs j = berghofe@18658: let berghofe@18658: val n = length vs; berghofe@18658: val Ts = map snd vs; berghofe@18658: val T = List.nth (Ts, i - j); berghofe@18658: val pT = NominalAtoms.mk_permT T berghofe@18658: in berghofe@18658: Const ("List.list.Cons", HOLogic.mk_prodT (T, T) --> pT --> pT) $ berghofe@22311: (HOLogic.pair_const T T $ Bound (l - j) $ fold_rev (mk_perm Ts) berghofe@18658: (map (mk_ind_perm i k p l vs) (j - 1 downto 0) @ berghofe@22311: map Bound (n - k - 1 downto n - k - p)) berghofe@22311: (Bound (i - j))) $ berghofe@18658: Const ("List.list.Nil", pT) berghofe@18658: end; berghofe@18658: berghofe@19710: fun mk_fresh i i' j k p l is vs _ _ = berghofe@18658: let berghofe@18658: val n = length vs; berghofe@18658: val Ts = map snd vs; berghofe@18658: val T = List.nth (Ts, n - i - 1 - j); berghofe@18658: val f = the (AList.lookup op = fresh_fs T); berghofe@18658: val U = List.nth (Ts, n - i' - 1); berghofe@18658: val S = HOLogic.mk_setT T; berghofe@19710: val prms' = map Bound (n - k downto n - k - p + 1); berghofe@18658: val prms = map (mk_ind_perm (n - i) k p (n - l) (("a", T) :: vs)) berghofe@19710: (j - 1 downto 0) @ prms'; berghofe@19710: val bs = rev (List.mapPartial berghofe@19710: (fn ((_, T'), i) => if T = T' then SOME (Bound i) else NONE) berghofe@19710: (List.take (vs, n - k - p - 1) ~~ (1 upto n - k - p - 1))); berghofe@19710: val ts = map (fn i => berghofe@19710: Const ("Nominal.supp", List.nth (Ts, n - i - 1) --> S) $ berghofe@22311: fold_rev (mk_perm (T :: Ts)) prms' (Bound (n - i))) is berghofe@18658: in berghofe@18658: HOLogic.mk_Trueprop (Const ("Ex", (T --> HOLogic.boolT) --> HOLogic.boolT) $ berghofe@18658: Abs ("a", T, HOLogic.Not $ berghofe@18658: (Const ("op :", T --> S --> HOLogic.boolT) $ Bound 0 $ berghofe@19710: (foldr (fn (t, u) => Const ("insert", T --> S --> S) $ t $ u) berghofe@19710: (foldl (fn (t, u) => Const ("op Un", S --> S --> S) $ u $ t) berghofe@19710: (f $ Bound (n - k - p)) berghofe@19710: (Const ("Nominal.supp", U --> S) $ berghofe@22311: fold_rev (mk_perm (T :: Ts)) prms (Bound (n - i')) :: ts)) berghofe@22311: (fold_rev (mk_perm (T :: Ts)) prms (Bound (n - i - j)) :: bs))))) berghofe@18658: end; urbanc@18104: berghofe@18658: fun mk_fresh_constr is p vs _ concl = berghofe@18658: let berghofe@18658: val n = length vs; berghofe@18658: val Ts = map snd vs; berghofe@18658: val _ $ (_ $ _ $ t) = concl; berghofe@18658: val c = head_of t; berghofe@18658: val T = body_type (fastype_of c); berghofe@18658: val k = foldr op + 0 (map (fn (_, i) => i + 1) is); berghofe@18658: val ps = map Bound (n - k - 1 downto n - k - p); berghofe@18658: val (_, _, ts, us) = foldl (fn ((_, i), (m, n, ts, us)) => berghofe@18658: (m - i - 1, n - i, berghofe@22311: ts @ map Bound (n downto n - i + 1) @ [fold_rev (mk_perm Ts) berghofe@22311: (map (mk_ind_perm m k p n vs) (i - 1 downto 0) @ ps) berghofe@22311: (Bound (m - i))], berghofe@22311: us @ map (fn j => fold_rev (mk_perm Ts) ps (Bound j)) (m downto m - i))) berghofe@18658: (n - 1, n - k - p - 2, [], []) is berghofe@18658: in berghofe@18658: HOLogic.mk_Trueprop (HOLogic.eq_const T $ list_comb (c, ts) $ list_comb (c, us)) berghofe@18658: end; berghofe@18658: berghofe@18658: val abs_fun_finite_supp = PureThy.get_thm thy9 (Name "abs_fun_finite_supp"); berghofe@18658: berghofe@18658: val at_finite_select = PureThy.get_thm thy9 (Name "at_finite_select"); berghofe@18658: berghofe@18658: val induct_aux_lemmas = List.concat (map (fn Type (s, _) => berghofe@18658: [PureThy.get_thm thy9 (Name ("pt_" ^ Sign.base_name s ^ "_inst")), berghofe@18658: PureThy.get_thm thy9 (Name ("fs_" ^ Sign.base_name s ^ "1")), berghofe@18658: PureThy.get_thm thy9 (Name ("at_" ^ Sign.base_name s ^ "_inst"))]) dt_atomTs); berghofe@18658: berghofe@18658: val induct_aux_lemmas' = map (fn Type (s, _) => berghofe@18658: PureThy.get_thm thy9 (Name ("pt_" ^ Sign.base_name s ^ "2")) RS sym) dt_atomTs; berghofe@18658: berghofe@19710: val fresh_aux = PureThy.get_thms thy9 (Name "fresh_aux"); berghofe@19710: berghofe@19710: (********************************************************************** berghofe@19710: The subgoals occurring in the proof of induct_aux have the berghofe@19710: following parameters: berghofe@19710: berghofe@19710: x_1 ... x_k p_1 ... p_m z b_1 ... b_n berghofe@19710: berghofe@19710: where berghofe@19710: berghofe@19710: x_i : constructor arguments (introduced by weak induction rule) berghofe@19710: p_i : permutations (one for each atom type in the data type) berghofe@19710: z : freshness context berghofe@19710: b_i : newly introduced names for binders (sufficiently fresh) berghofe@19710: ***********************************************************************) berghofe@19710: berghofe@19710: val _ = warning "proving strong induction theorem ..."; berghofe@19710: wenzelm@20046: val induct_aux = Goal.prove_global thy9 [] ind_prems' ind_concl' berghofe@18658: (fn prems => EVERY berghofe@18658: ([mk_subgoal 1 (K (K (K aux_ind_concl))), berghofe@18658: indtac dt_induct tnames 1] @ berghofe@18658: List.concat (map (fn ((_, (_, _, constrs)), (_, constrs')) => berghofe@18658: List.concat (map (fn ((cname, cargs), is) => berghofe@18658: [simp_tac (HOL_basic_ss addsimps List.concat perm_simps') 1, berghofe@18658: REPEAT (rtac allI 1)] @ berghofe@18658: List.concat (map berghofe@18658: (fn ((_, 0), _) => [] berghofe@18658: | ((i, j), k) => List.concat (map (fn j' => berghofe@18658: let berghofe@18658: val DtType (tname, _) = List.nth (cargs, i + j'); berghofe@18658: val atom = Sign.base_name tname berghofe@18658: in berghofe@18658: [mk_subgoal 1 (mk_fresh i (i + j) j' berghofe@18658: (length cargs) (length dt_atomTs) berghofe@19710: (length cargs + length dt_atomTs + 1 + i - k) berghofe@19710: (List.mapPartial (fn (i', j) => berghofe@19710: if i = i' then NONE else SOME (i' + j)) is)), berghofe@18658: rtac at_finite_select 1, berghofe@18658: rtac (PureThy.get_thm thy9 (Name ("at_" ^ atom ^ "_inst"))) 1, berghofe@18658: asm_full_simp_tac (simpset_of thy9 addsimps berghofe@18658: [PureThy.get_thm thy9 (Name ("fs_" ^ atom ^ "1"))]) 1, berghofe@18658: resolve_tac prems 1, berghofe@18658: etac exE 1, berghofe@18658: asm_full_simp_tac (simpset_of thy9 addsimps berghofe@18658: [symmetric fresh_def]) 1] berghofe@18658: end) (0 upto j - 1))) (is ~~ (0 upto length is - 1))) @ berghofe@18658: (if exists (not o equal 0 o snd) is then berghofe@18658: [mk_subgoal 1 (mk_fresh_constr is (length dt_atomTs)), berghofe@18658: asm_full_simp_tac (simpset_of thy9 addsimps berghofe@18658: (List.concat inject_thms @ berghofe@18658: alpha @ abs_perm @ abs_fresh @ [abs_fun_finite_supp] @ berghofe@18658: induct_aux_lemmas)) 1, berghofe@18658: dtac sym 1, berghofe@19710: asm_full_simp_tac (simpset_of thy9) 1, berghofe@18658: REPEAT (etac conjE 1)] berghofe@18658: else berghofe@18658: []) @ berghofe@18658: [(resolve_tac prems THEN_ALL_NEW berghofe@19710: (atac ORELSE' berghofe@19710: SUBGOAL (fn (t, i) => case Logic.strip_assums_concl t of berghofe@19710: _ $ (Const ("Nominal.fresh", _) $ _ $ _) => berghofe@19710: asm_simp_tac (simpset_of thy9 addsimps fresh_aux) i berghofe@19710: | _ => berghofe@19710: asm_simp_tac (simpset_of thy9 berghofe@19710: addsimps induct_aux_lemmas' berghofe@19710: addsimprocs [perm_simproc]) i))) 1]) berghofe@18658: (constrs ~~ constrs'))) (descr'' ~~ ndescr)) @ berghofe@18658: [REPEAT (eresolve_tac [conjE, allE_Nil] 1), berghofe@18658: REPEAT (etac allE 1), wenzelm@20046: REPEAT (TRY (rtac conjI 1) THEN asm_full_simp_tac (simpset_of thy9) 1)])); berghofe@18658: berghofe@18658: val induct_aux' = Thm.instantiate ([], berghofe@18658: map (fn (s, T) => berghofe@18658: let val pT = TVar (("'n", 0), HOLogic.typeS) --> T --> HOLogic.boolT berghofe@18658: in (cterm_of thy9 (Var ((s, 0), pT)), cterm_of thy9 (Free (s, pT))) end) berghofe@18658: (pnames ~~ recTs) @ berghofe@18658: map (fn (_, f) => berghofe@18658: let val f' = Logic.varify f berghofe@18658: in (cterm_of thy9 f', berghofe@19494: cterm_of thy9 (Const ("Nominal.supp", fastype_of f'))) berghofe@18658: end) fresh_fs) induct_aux; berghofe@18658: wenzelm@20046: val induct = Goal.prove_global thy9 [] ind_prems ind_concl berghofe@18658: (fn prems => EVERY berghofe@18658: [rtac induct_aux' 1, berghofe@18658: REPEAT (resolve_tac induct_aux_lemmas 1), berghofe@18658: REPEAT ((resolve_tac prems THEN_ALL_NEW wenzelm@20046: (etac meta_spec ORELSE' full_simp_tac (HOL_basic_ss addsimps [fresh_def]))) 1)]) berghofe@18658: berghofe@18658: val (_, thy10) = thy9 |> wenzelm@24712: Sign.add_path big_name |> berghofe@18658: PureThy.add_thms [(("induct'", induct_aux), [])] ||>> berghofe@18658: PureThy.add_thms [(("induct", induct), [case_names_induct])] ||>> berghofe@20397: PureThy.add_thmss [(("inducts", projections induct), [case_names_induct])]; berghofe@18658: berghofe@19322: (**** recursion combinator ****) berghofe@19251: berghofe@19322: val _ = warning "defining recursion combinator ..."; berghofe@19251: berghofe@19251: val used = foldr add_typ_tfree_names [] recTs; berghofe@19251: berghofe@19985: val (rec_result_Ts', rec_fn_Ts') = DatatypeProp.make_primrec_Ts descr' sorts' used; berghofe@19985: wenzelm@21365: val rec_sort = if null dt_atomTs then HOLogic.typeS else berghofe@19985: let val names = map (Sign.base_name o fst o dest_Type) dt_atomTs berghofe@19985: in Sign.certify_sort thy10 (map (Sign.intern_class thy10) berghofe@19985: (map (fn s => "pt_" ^ s) names @ berghofe@19985: List.concat (map (fn s => List.mapPartial (fn s' => berghofe@19985: if s = s' then NONE berghofe@19985: else SOME ("cp_" ^ s ^ "_" ^ s')) names) names))) berghofe@19985: end; berghofe@19985: berghofe@19985: val rec_result_Ts = map (fn TFree (s, _) => TFree (s, rec_sort)) rec_result_Ts'; berghofe@19985: val rec_fn_Ts = map (typ_subst_atomic (rec_result_Ts' ~~ rec_result_Ts)) rec_fn_Ts'; berghofe@19251: berghofe@21021: val rec_set_Ts = map (fn (T1, T2) => berghofe@21021: rec_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts); berghofe@19251: berghofe@19322: 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'')); wenzelm@22578: val rec_set_names = map (Sign.full_name thy10) rec_set_names'; berghofe@19251: berghofe@19322: val rec_fns = map (uncurry (mk_Free "f")) berghofe@19322: (rec_fn_Ts ~~ (1 upto (length rec_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@19322: val rec_sets = map (fn c => list_comb (Const c, rec_fns)) berghofe@19322: (rec_set_names ~~ rec_set_Ts); berghofe@19251: berghofe@19322: (* introduction rules for graph of recursion function *) berghofe@19251: berghofe@20145: val rec_preds = map (fn (a, T) => berghofe@20145: Free (a, T --> HOLogic.boolT)) (pnames ~~ rec_result_Ts); berghofe@20145: berghofe@20267: fun mk_fresh3 rs [] = [] berghofe@20267: | mk_fresh3 rs ((p as (ys, z)) :: yss) = List.concat (map (fn y as (_, T) => berghofe@20267: List.mapPartial (fn ((_, (_, x)), r as (_, U)) => if z = x then NONE berghofe@20267: else SOME (HOLogic.mk_Trueprop berghofe@20267: (Const ("Nominal.fresh", T --> U --> HOLogic.boolT) $ Free y $ Free r))) berghofe@20267: rs) ys) @ berghofe@20267: mk_fresh3 rs yss; berghofe@20267: berghofe@21088: (* FIXME: avoid collisions with other variable names? *) berghofe@21088: val rec_ctxt = Free ("z", fsT'); berghofe@21088: berghofe@20397: fun make_rec_intr T p rec_set ((rec_intr_ts, rec_prems, rec_prems', berghofe@20397: rec_eq_prems, l), ((cname, cargs), idxs)) = berghofe@19251: let berghofe@19251: val Ts = map (typ_of_dtyp descr'' sorts') cargs; berghofe@19851: val frees = map (fn i => "x" ^ string_of_int i) (1 upto length Ts) ~~ Ts; berghofe@19851: val frees' = partition_cargs idxs frees; berghofe@21088: val binders = List.concat (map fst frees'); berghofe@20411: val atomTs = distinct op = (maps (map snd o fst) frees'); berghofe@19851: val recs = List.mapPartial berghofe@20145: (fn ((_, DtRec i), p) => SOME (i, p) | _ => NONE) berghofe@19851: (partition_cargs idxs cargs ~~ frees'); berghofe@19851: val frees'' = map (fn i => "y" ^ string_of_int i) (1 upto length recs) ~~ berghofe@19851: map (fn (i, _) => List.nth (rec_result_Ts, i)) recs; berghofe@20145: val prems1 = map (fn ((i, (_, x)), y) => HOLogic.mk_Trueprop berghofe@21021: (List.nth (rec_sets', i) $ Free x $ Free y)) (recs ~~ frees''); berghofe@20145: val prems2 = berghofe@20145: map (fn f => map (fn p as (_, T) => HOLogic.mk_Trueprop berghofe@19851: (Const ("Nominal.fresh", T --> fastype_of f --> HOLogic.boolT) $ berghofe@21088: Free p $ f)) binders) rec_fns; berghofe@21088: val prems3 = mk_fresh1 [] binders @ mk_fresh2 [] frees'; berghofe@20145: val prems4 = map (fn ((i, _), y) => berghofe@20145: HOLogic.mk_Trueprop (List.nth (rec_preds, i) $ Free y)) (recs ~~ frees''); berghofe@20267: val prems5 = mk_fresh3 (recs ~~ frees'') frees'; berghofe@20411: val prems6 = maps (fn aT => map (fn y as (_, T) => HOLogic.mk_Trueprop berghofe@22274: (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $ berghofe@22274: (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ Free y))) berghofe@20411: frees'') atomTs; berghofe@21088: val prems7 = map (fn x as (_, T) => HOLogic.mk_Trueprop berghofe@21088: (Const ("Nominal.fresh", T --> fsT' --> HOLogic.boolT) $ berghofe@21088: Free x $ rec_ctxt)) binders; berghofe@20145: val result = list_comb (List.nth (rec_fns, l), map Free (frees @ frees'')); berghofe@20376: val result_freshs = map (fn p as (_, T) => berghofe@20145: Const ("Nominal.fresh", T --> fastype_of result --> HOLogic.boolT) $ berghofe@21088: Free p $ result) binders; berghofe@20145: val P = HOLogic.mk_Trueprop (p $ result) berghofe@20145: in berghofe@20145: (rec_intr_ts @ [Logic.list_implies (List.concat prems2 @ prems3 @ prems1, berghofe@21021: HOLogic.mk_Trueprop (rec_set $ berghofe@21021: list_comb (Const (cname, Ts ---> T), map Free frees) $ result))], berghofe@20145: rec_prems @ [list_all_free (frees @ frees'', Logic.list_implies (prems4, P))], berghofe@21054: rec_prems' @ map (fn fr => list_all_free (frees @ frees'', berghofe@21088: Logic.list_implies (List.nth (prems2, l) @ prems3 @ prems5 @ prems7 @ prems6 @ [P], berghofe@21054: HOLogic.mk_Trueprop fr))) result_freshs, berghofe@20397: rec_eq_prems @ [List.concat prems2 @ prems3], berghofe@20145: l + 1) berghofe@19251: end; berghofe@19251: berghofe@20397: val (rec_intr_ts, rec_prems, rec_prems', rec_eq_prems, _) = berghofe@20145: Library.foldl (fn (x, ((((d, d'), T), p), rec_set)) => berghofe@20145: Library.foldl (make_rec_intr T p rec_set) (x, #3 (snd d) ~~ snd d')) berghofe@21021: (([], [], [], [], 0), descr'' ~~ ndescr ~~ recTs ~~ rec_preds ~~ rec_sets'); berghofe@19251: wenzelm@21365: val ({intrs = rec_intrs, elims = rec_elims, raw_induct = rec_induct, ...}, thy11) = berghofe@21055: thy10 |> berghofe@19251: setmp InductivePackage.quiet_mode (!quiet_mode) wenzelm@24814: (InductivePackage.add_inductive_global {verbose = false, kind = Thm.internalK, wenzelm@24814: alt_name = big_rec_name, coind = false, no_elim = false, no_ind = false} berghofe@24746: (map (fn (s, T) => ((s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts)) berghofe@24746: (map dest_Free rec_fns) berghofe@22607: (map (fn x => (("", []), x)) rec_intr_ts) []) ||> berghofe@21055: PureThy.hide_thms true [NameSpace.append berghofe@21055: (Sign.full_name thy10 big_rec_name) "induct"]; berghofe@19251: berghofe@19985: (** equivariance **) berghofe@19985: berghofe@19985: val fresh_bij = PureThy.get_thms thy11 (Name "fresh_bij"); berghofe@19985: val perm_bij = PureThy.get_thms thy11 (Name "perm_bij"); berghofe@19985: berghofe@19985: val (rec_equiv_thms, rec_equiv_thms') = ListPair.unzip (map (fn aT => berghofe@19985: let berghofe@19985: val permT = mk_permT aT; berghofe@19985: val pi = Free ("pi", permT); berghofe@22311: val rec_fns_pi = map (mk_perm [] pi o uncurry (mk_Free "f")) berghofe@19985: (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts))); berghofe@19985: val rec_sets_pi = map (fn c => list_comb (Const c, rec_fns_pi)) berghofe@19985: (rec_set_names ~~ rec_set_Ts); berghofe@19985: val ps = map (fn ((((T, U), R), R'), i) => berghofe@19985: let berghofe@19985: val x = Free ("x" ^ string_of_int i, T); berghofe@19985: val y = Free ("y" ^ string_of_int i, U) berghofe@19985: in berghofe@22311: (R $ x $ y, R' $ mk_perm [] pi x $ mk_perm [] pi y) berghofe@19985: end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ rec_sets_pi ~~ (1 upto length recTs)); berghofe@19985: val ths = map (fn th => standard (th RS mp)) (split_conj_thm wenzelm@20046: (Goal.prove_global thy11 [] [] berghofe@19985: (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps))) berghofe@19985: (fn _ => rtac rec_induct 1 THEN REPEAT berghofe@22433: (NominalPermeq.perm_simp_tac (HOL_basic_ss addsimps flat perm_simps') 1 THEN berghofe@19985: (resolve_tac rec_intrs THEN_ALL_NEW berghofe@19985: asm_simp_tac (HOL_ss addsimps (fresh_bij @ perm_bij))) 1)))) wenzelm@20046: val ths' = map (fn ((P, Q), th) => wenzelm@20046: Goal.prove_global thy11 [] [] berghofe@19985: (Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P)) berghofe@19985: (fn _ => dtac (Thm.instantiate ([], berghofe@19985: [(cterm_of thy11 (Var (("pi", 0), permT)), berghofe@19985: cterm_of thy11 (Const ("List.rev", permT --> permT) $ pi))]) th) 1 THEN wenzelm@20046: NominalPermeq.perm_simp_tac HOL_ss 1)) (ps ~~ ths) berghofe@19985: in (ths, ths') end) dt_atomTs); berghofe@19985: berghofe@19985: (** finite support **) berghofe@19985: berghofe@19985: val rec_fin_supp_thms = map (fn aT => berghofe@19985: let berghofe@19985: val name = Sign.base_name (fst (dest_Type aT)); berghofe@19985: val fs_name = PureThy.get_thm thy11 (Name ("fs_" ^ name ^ "1")); berghofe@19985: val aset = HOLogic.mk_setT aT; berghofe@22274: val finite = Const ("Finite_Set.finite", aset --> HOLogic.boolT); berghofe@22274: val fins = map (fn (f, T) => HOLogic.mk_Trueprop berghofe@22274: (finite $ (Const ("Nominal.supp", T --> aset) $ f))) berghofe@19985: (rec_fns ~~ rec_fn_Ts) berghofe@19985: in berghofe@19985: map (fn th => standard (th RS mp)) (split_conj_thm wenzelm@20046: (Goal.prove_global thy11 [] fins berghofe@19985: (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj berghofe@19985: (map (fn (((T, U), R), i) => berghofe@19985: let berghofe@19985: val x = Free ("x" ^ string_of_int i, T); berghofe@19985: val y = Free ("y" ^ string_of_int i, U) berghofe@19985: in berghofe@21021: HOLogic.mk_imp (R $ x $ y, berghofe@22274: finite $ (Const ("Nominal.supp", U --> aset) $ y)) berghofe@19985: end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ (1 upto length recTs))))) berghofe@19985: (fn fins => (rtac rec_induct THEN_ALL_NEW cut_facts_tac fins) 1 THEN REPEAT berghofe@19985: (NominalPermeq.finite_guess_tac (HOL_ss addsimps [fs_name]) 1)))) berghofe@19985: end) dt_atomTs; berghofe@19985: berghofe@20376: (** freshness **) berghofe@20376: berghofe@20376: val perm_fresh_fresh = PureThy.get_thms thy11 (Name "perm_fresh_fresh"); berghofe@20376: val perm_swap = PureThy.get_thms thy11 (Name "perm_swap"); berghofe@20145: berghofe@20376: val finite_premss = map (fn aT => berghofe@22274: map (fn (f, T) => HOLogic.mk_Trueprop berghofe@22274: (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $ berghofe@22274: (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ f))) berghofe@20376: (rec_fns ~~ rec_fn_Ts)) dt_atomTs; berghofe@20376: berghofe@20376: val rec_fresh_thms = map (fn ((aT, eqvt_ths), finite_prems) => berghofe@20376: let berghofe@20376: val name = Sign.base_name (fst (dest_Type aT)); berghofe@20376: val fs_name = PureThy.get_thm thy11 (Name ("fs_" ^ name ^ "1")); berghofe@20376: val a = Free ("a", aT); berghofe@20376: val freshs = map (fn (f, fT) => HOLogic.mk_Trueprop berghofe@20376: (Const ("Nominal.fresh", aT --> fT --> HOLogic.boolT) $ a $ f)) berghofe@20376: (rec_fns ~~ rec_fn_Ts) berghofe@20376: in berghofe@20376: map (fn (((T, U), R), eqvt_th) => berghofe@20376: let berghofe@20376: val x = Free ("x", T); berghofe@20376: val y = Free ("y", U); berghofe@20376: val y' = Free ("y'", U) berghofe@20376: in wenzelm@21516: standard (Goal.prove (ProofContext.init thy11) [] (finite_prems @ berghofe@21021: [HOLogic.mk_Trueprop (R $ x $ y), berghofe@20376: HOLogic.mk_Trueprop (HOLogic.mk_all ("y'", U, berghofe@21021: HOLogic.mk_imp (R $ x $ y', HOLogic.mk_eq (y', y)))), berghofe@20376: HOLogic.mk_Trueprop (Const ("Nominal.fresh", berghofe@20376: aT --> T --> HOLogic.boolT) $ a $ x)] @ berghofe@20376: freshs) berghofe@20376: (HOLogic.mk_Trueprop (Const ("Nominal.fresh", berghofe@20376: aT --> U --> HOLogic.boolT) $ a $ y)) berghofe@20376: (fn {prems, context} => berghofe@20376: let berghofe@20376: val (finite_prems, rec_prem :: unique_prem :: berghofe@20376: fresh_prems) = chop (length finite_prems) prems; berghofe@20376: val unique_prem' = unique_prem RS spec RS mp; berghofe@20376: val unique = [unique_prem', unique_prem' RS sym] MRS trans; berghofe@20376: val _ $ (_ $ (_ $ S $ _)) $ _ = prop_of supports_fresh; berghofe@20376: val tuple = foldr1 HOLogic.mk_prod (x :: rec_fns) berghofe@20376: in EVERY berghofe@20376: [rtac (Drule.cterm_instantiate berghofe@20376: [(cterm_of thy11 S, berghofe@20376: cterm_of thy11 (Const ("Nominal.supp", berghofe@20376: fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))] berghofe@20376: supports_fresh) 1, berghofe@20376: simp_tac (HOL_basic_ss addsimps berghofe@20376: [supports_def, symmetric fresh_def, fresh_prod]) 1, berghofe@20376: REPEAT_DETERM (resolve_tac [allI, impI] 1), berghofe@20376: REPEAT_DETERM (etac conjE 1), berghofe@20376: rtac unique 1, berghofe@20376: SUBPROOF (fn {prems = prems', params = [a, b], ...} => EVERY berghofe@20376: [cut_facts_tac [rec_prem] 1, berghofe@20376: rtac (Thm.instantiate ([], berghofe@20376: [(cterm_of thy11 (Var (("pi", 0), mk_permT aT)), berghofe@20376: cterm_of thy11 (perm_of_pair (term_of a, term_of b)))]) eqvt_th) 1, berghofe@20376: asm_simp_tac (HOL_ss addsimps berghofe@20376: (prems' @ perm_swap @ perm_fresh_fresh)) 1]) context 1, berghofe@20376: rtac rec_prem 1, berghofe@20376: simp_tac (HOL_ss addsimps (fs_name :: berghofe@20376: supp_prod :: finite_Un :: finite_prems)) 1, berghofe@20376: simp_tac (HOL_ss addsimps (symmetric fresh_def :: berghofe@20376: fresh_prod :: fresh_prems)) 1] berghofe@20376: end)) berghofe@20376: end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ eqvt_ths) berghofe@20376: end) (dt_atomTs ~~ rec_equiv_thms' ~~ finite_premss); berghofe@20376: berghofe@20376: (** uniqueness **) berghofe@20376: urbanc@21377: val exists_fresh' = PureThy.get_thms thy11 (Name "exists_fresh'"); berghofe@20376: val fs_atoms = map (fn Type (s, _) => PureThy.get_thm thy11 berghofe@20376: (Name ("fs_" ^ Sign.base_name s ^ "1"))) dt_atomTs; berghofe@20376: val fresh_atm = PureThy.get_thms thy11 (Name "fresh_atm"); berghofe@20376: val calc_atm = PureThy.get_thms thy11 (Name "calc_atm"); berghofe@20376: val fresh_left = PureThy.get_thms thy11 (Name "fresh_left"); berghofe@20145: berghofe@21088: val fun_tuple = foldr1 HOLogic.mk_prod (rec_ctxt :: rec_fns); berghofe@20145: val fun_tupleT = fastype_of fun_tuple; berghofe@20145: val rec_unique_frees = berghofe@20145: DatatypeProp.indexify_names (replicate (length recTs) "x") ~~ recTs; berghofe@20397: val rec_unique_frees'' = map (fn (s, T) => (s ^ "'", T)) rec_unique_frees; berghofe@20267: val rec_unique_frees' = berghofe@20267: DatatypeProp.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts; berghofe@21021: val rec_unique_concls = map (fn ((x, U), R) => berghofe@20145: Const ("Ex1", (U --> HOLogic.boolT) --> HOLogic.boolT) $ berghofe@21021: Abs ("y", U, R $ Free x $ Bound 0)) berghofe@20145: (rec_unique_frees ~~ rec_result_Ts ~~ rec_sets); berghofe@20376: berghofe@20145: val induct_aux_rec = Drule.cterm_instantiate berghofe@20145: (map (pairself (cterm_of thy11)) berghofe@20145: (map (fn (aT, f) => (Logic.varify f, Abs ("z", HOLogic.unitT, berghofe@20145: Const ("Nominal.supp", fun_tupleT --> HOLogic.mk_setT aT) $ fun_tuple))) berghofe@20145: fresh_fs @ berghofe@20145: map (fn (((P, T), (x, U)), Q) => berghofe@20145: (Var ((P, 0), HOLogic.unitT --> Logic.varifyT T --> HOLogic.boolT), berghofe@20145: Abs ("z", HOLogic.unitT, absfree (x, U, Q)))) berghofe@20145: (pnames ~~ recTs ~~ rec_unique_frees ~~ rec_unique_concls) @ berghofe@20145: map (fn (s, T) => (Var ((s, 0), Logic.varifyT T), Free (s, T))) berghofe@20145: rec_unique_frees)) induct_aux; berghofe@20145: berghofe@20376: fun obtain_fresh_name vs ths rec_fin_supp T (freshs1, freshs2, ctxt) = berghofe@20376: let berghofe@20376: val p = foldr1 HOLogic.mk_prod (vs @ freshs1); berghofe@20376: val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop berghofe@20376: (HOLogic.exists_const T $ Abs ("x", T, berghofe@20376: Const ("Nominal.fresh", T --> fastype_of p --> HOLogic.boolT) $ berghofe@20376: Bound 0 $ p))) berghofe@20376: (fn _ => EVERY berghofe@20376: [cut_facts_tac ths 1, berghofe@20376: REPEAT_DETERM (dresolve_tac (the (AList.lookup op = rec_fin_supp T)) 1), urbanc@21377: resolve_tac exists_fresh' 1, berghofe@20376: asm_simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]); berghofe@20376: val (([cx], ths), ctxt') = Obtain.result berghofe@20376: (fn _ => EVERY berghofe@20376: [etac exE 1, berghofe@20376: full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1, berghofe@20376: REPEAT (etac conjE 1)]) berghofe@20376: [ex] ctxt berghofe@20376: in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end; berghofe@20376: berghofe@21088: val finite_ctxt_prems = map (fn aT => berghofe@22274: HOLogic.mk_Trueprop berghofe@22274: (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $ berghofe@22274: (Const ("Nominal.supp", fsT' --> HOLogic.mk_setT aT) $ rec_ctxt))) dt_atomTs; berghofe@21088: berghofe@20397: val rec_unique_thms = split_conj_thm (Goal.prove wenzelm@21516: (ProofContext.init thy11) (map fst rec_unique_frees) berghofe@21088: (List.concat finite_premss @ finite_ctxt_prems @ rec_prems @ rec_prems') berghofe@20145: (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj rec_unique_concls)) berghofe@20376: (fn {prems, context} => berghofe@20267: let berghofe@20267: val k = length rec_fns; berghofe@20376: val (finite_thss, ths1) = fold_map (fn T => fn xs => berghofe@20376: apfst (pair T) (chop k xs)) dt_atomTs prems; berghofe@21088: val (finite_ctxt_ths, ths2) = chop (length dt_atomTs) ths1; berghofe@21088: val (P_ind_ths, fcbs) = chop k ths2; berghofe@20267: val P_ths = map (fn th => th RS mp) (split_conj_thm berghofe@20376: (Goal.prove context berghofe@20397: (map fst (rec_unique_frees'' @ rec_unique_frees')) [] berghofe@20267: (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj berghofe@21021: (map (fn (((x, y), S), P) => HOLogic.mk_imp berghofe@21021: (S $ Free x $ Free y, P $ (Free y))) berghofe@20397: (rec_unique_frees'' ~~ rec_unique_frees' ~~ rec_sets ~~ rec_preds)))) berghofe@20267: (fn _ => berghofe@20267: rtac rec_induct 1 THEN berghofe@20376: REPEAT ((resolve_tac P_ind_ths THEN_ALL_NEW assume_tac) 1)))); berghofe@20376: val rec_fin_supp_thms' = map berghofe@20376: (fn (ths, (T, fin_ths)) => (T, map (curry op MRS fin_ths) ths)) berghofe@20376: (rec_fin_supp_thms ~~ finite_thss); berghofe@20267: in EVERY berghofe@20267: ([rtac induct_aux_rec 1] @ berghofe@21088: maps (fn ((_, finite_ths), finite_th) => berghofe@21088: [cut_facts_tac (finite_th :: finite_ths) 1, berghofe@21088: asm_simp_tac (HOL_ss addsimps [supp_prod, finite_Un]) 1]) berghofe@21088: (finite_thss ~~ finite_ctxt_ths) @ berghofe@20397: maps (fn ((_, idxss), elim) => maps (fn idxs => berghofe@20397: [full_simp_tac (HOL_ss addsimps [symmetric fresh_def, supp_prod, Un_iff]) 1, berghofe@20397: REPEAT_DETERM (eresolve_tac [conjE, ex1E] 1), berghofe@20397: rtac ex1I 1, berghofe@20397: (resolve_tac rec_intrs THEN_ALL_NEW atac) 1, berghofe@20397: rotate_tac ~1 1, berghofe@20397: ((DETERM o etac elim) THEN_ALL_NEW full_simp_tac berghofe@21021: (HOL_ss addsimps List.concat distinct_thms)) 1] @ berghofe@21021: (if null idxs then [] else [hyp_subst_tac 1, berghofe@20397: SUBPROOF (fn {asms, concl, prems = prems', params, context = context', ...} => berghofe@20376: let berghofe@21021: val SOME prem = find_first (can (HOLogic.dest_eq o berghofe@21021: HOLogic.dest_Trueprop o prop_of)) prems'; berghofe@20376: val _ $ (_ $ lhs $ rhs) = prop_of prem; berghofe@20376: val _ $ (_ $ lhs' $ rhs') = term_of concl; berghofe@20376: val rT = fastype_of lhs'; berghofe@20376: val (c, cargsl) = strip_comb lhs; berghofe@20376: val cargsl' = partition_cargs idxs cargsl; berghofe@20376: val boundsl = List.concat (map fst cargsl'); berghofe@20376: val (_, cargsr) = strip_comb rhs; berghofe@20376: val cargsr' = partition_cargs idxs cargsr; berghofe@20376: val boundsr = List.concat (map fst cargsr'); berghofe@20376: val (params1, _ :: params2) = berghofe@20376: chop (length params div 2) (map term_of params); berghofe@20376: val params' = params1 @ params2; berghofe@20376: val rec_prems = filter (fn th => case prop_of th of berghofe@21021: _ $ (S $ _ $ _) => S mem rec_sets | _ => false) prems'; berghofe@20376: val fresh_prems = filter (fn th => case prop_of th of berghofe@20376: _ $ (Const ("Nominal.fresh", _) $ _ $ _) => true berghofe@20376: | _ $ (Const ("Not", _) $ _) => true berghofe@20376: | _ => false) prems'; berghofe@20376: val Ts = map fastype_of boundsl; berghofe@20376: berghofe@20376: val _ = warning "step 1: obtaining fresh names"; berghofe@20376: val (freshs1, freshs2, context'') = fold berghofe@21088: (obtain_fresh_name (rec_ctxt :: rec_fns @ params') berghofe@21088: (List.concat (map snd finite_thss) @ berghofe@21088: finite_ctxt_ths @ rec_prems) berghofe@20376: rec_fin_supp_thms') berghofe@20376: Ts ([], [], context'); berghofe@20376: val pi1 = map perm_of_pair (boundsl ~~ freshs1); berghofe@20376: val rpi1 = rev pi1; berghofe@20376: val pi2 = map perm_of_pair (boundsr ~~ freshs1); berghofe@21073: val rpi2 = rev pi2; berghofe@20376: berghofe@20376: val fresh_prems' = mk_not_sym fresh_prems; berghofe@20376: val freshs2' = mk_not_sym freshs2; berghofe@20376: berghofe@20376: (** as, bs, cs # K as ts, K bs us **) berghofe@20376: val _ = warning "step 2: as, bs, cs # K as ts, K bs us"; berghofe@20376: val prove_fresh_ss = HOL_ss addsimps berghofe@20376: (finite_Diff :: List.concat fresh_thms @ berghofe@20376: fs_atoms @ abs_fresh @ abs_supp @ fresh_atm); berghofe@20376: (* FIXME: avoid asm_full_simp_tac ? *) berghofe@20376: fun prove_fresh ths y x = Goal.prove context'' [] [] berghofe@20376: (HOLogic.mk_Trueprop (Const ("Nominal.fresh", berghofe@20376: fastype_of x --> fastype_of y --> HOLogic.boolT) $ x $ y)) berghofe@20376: (fn _ => cut_facts_tac ths 1 THEN asm_full_simp_tac prove_fresh_ss 1); berghofe@20376: val constr_fresh_thms = berghofe@20376: map (prove_fresh fresh_prems lhs) boundsl @ berghofe@20376: map (prove_fresh fresh_prems rhs) boundsr @ berghofe@20376: map (prove_fresh freshs2 lhs) freshs1 @ berghofe@20376: map (prove_fresh freshs2 rhs) freshs1; berghofe@20376: berghofe@20376: (** pi1 o (K as ts) = pi2 o (K bs us) **) berghofe@20376: val _ = warning "step 3: pi1 o (K as ts) = pi2 o (K bs us)"; berghofe@20376: val pi1_pi2_eq = Goal.prove context'' [] [] berghofe@20376: (HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@22311: (fold_rev (mk_perm []) pi1 lhs, fold_rev (mk_perm []) pi2 rhs))) berghofe@20376: (fn _ => EVERY berghofe@20376: [cut_facts_tac constr_fresh_thms 1, berghofe@20376: asm_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh) 1, berghofe@20376: rtac prem 1]); berghofe@20376: berghofe@20376: (** pi1 o ts = pi2 o us **) berghofe@20376: val _ = warning "step 4: pi1 o ts = pi2 o us"; berghofe@20376: val pi1_pi2_eqs = map (fn (t, u) => berghofe@20376: Goal.prove context'' [] [] berghofe@20376: (HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@22311: (fold_rev (mk_perm []) pi1 t, fold_rev (mk_perm []) pi2 u))) berghofe@20376: (fn _ => EVERY berghofe@20376: [cut_facts_tac [pi1_pi2_eq] 1, berghofe@20376: asm_full_simp_tac (HOL_ss addsimps berghofe@20376: (calc_atm @ List.concat perm_simps' @ berghofe@20376: fresh_prems' @ freshs2' @ abs_perm @ berghofe@20376: alpha @ List.concat inject_thms)) 1])) berghofe@20376: (map snd cargsl' ~~ map snd cargsr'); berghofe@20376: berghofe@20376: (** pi1^-1 o pi2 o us = ts **) berghofe@20376: val _ = warning "step 5: pi1^-1 o pi2 o us = ts"; berghofe@20376: val rpi1_pi2_eqs = map (fn ((t, u), eq) => berghofe@20376: Goal.prove context'' [] [] berghofe@20376: (HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@22311: (fold_rev (mk_perm []) (rpi1 @ pi2) u, t))) berghofe@20376: (fn _ => simp_tac (HOL_ss addsimps berghofe@20376: ((eq RS sym) :: perm_swap)) 1)) berghofe@20376: (map snd cargsl' ~~ map snd cargsr' ~~ pi1_pi2_eqs); berghofe@20376: berghofe@20376: val (rec_prems1, rec_prems2) = berghofe@20376: chop (length rec_prems div 2) rec_prems; berghofe@20376: berghofe@20376: (** (ts, pi1^-1 o pi2 o vs) in rec_set **) berghofe@20376: val _ = warning "step 6: (ts, pi1^-1 o pi2 o vs) in rec_set"; berghofe@20376: val rec_prems' = map (fn th => berghofe@20376: let berghofe@21021: val _ $ (S $ x $ y) = prop_of th; berghofe@20376: val k = find_index (equal S) rec_sets; berghofe@20376: val pi = rpi1 @ pi2; berghofe@22311: fun mk_pi z = fold_rev (mk_perm []) pi z; berghofe@20376: fun eqvt_tac p = berghofe@20376: let berghofe@20376: val U as Type (_, [Type (_, [T, _])]) = fastype_of p; berghofe@20376: val l = find_index (equal T) dt_atomTs; berghofe@20376: val th = List.nth (List.nth (rec_equiv_thms', l), k); berghofe@20376: val th' = Thm.instantiate ([], berghofe@20376: [(cterm_of thy11 (Var (("pi", 0), U)), berghofe@20376: cterm_of thy11 p)]) th; berghofe@20376: in rtac th' 1 end; berghofe@20376: val th' = Goal.prove context'' [] [] berghofe@21021: (HOLogic.mk_Trueprop (S $ mk_pi x $ mk_pi y)) berghofe@20376: (fn _ => EVERY berghofe@20376: (map eqvt_tac pi @ berghofe@20376: [simp_tac (HOL_ss addsimps (fresh_prems' @ freshs2' @ berghofe@20376: perm_swap @ perm_fresh_fresh)) 1, berghofe@20376: rtac th 1])) berghofe@20376: in berghofe@20376: Simplifier.simplify berghofe@20376: (HOL_basic_ss addsimps rpi1_pi2_eqs) th' berghofe@20376: end) rec_prems2; berghofe@20376: berghofe@20376: val ihs = filter (fn th => case prop_of th of berghofe@20376: _ $ (Const ("All", _) $ _) => true | _ => false) prems'; berghofe@20376: berghofe@21073: (** pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs **) berghofe@21073: val _ = warning "step 7: pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs"; berghofe@21073: val rec_eqns = map (fn (th, ih) => berghofe@20376: let berghofe@20376: val th' = th RS (ih RS spec RS mp) RS sym; berghofe@20376: val _ $ (_ $ lhs $ rhs) = prop_of th'; berghofe@20376: fun strip_perm (_ $ _ $ t) = strip_perm t berghofe@20376: | strip_perm t = t; berghofe@20376: in berghofe@21073: Goal.prove context'' [] [] berghofe@20376: (HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@22311: (fold_rev (mk_perm []) pi1 lhs, berghofe@22311: fold_rev (mk_perm []) pi2 (strip_perm rhs)))) berghofe@20376: (fn _ => simp_tac (HOL_basic_ss addsimps berghofe@21073: (th' :: perm_swap)) 1) berghofe@21073: end) (rec_prems' ~~ ihs); berghofe@20376: berghofe@21073: (** as # rs **) berghofe@21073: val _ = warning "step 8: as # rs"; berghofe@21073: val rec_freshs = List.concat berghofe@21073: (map (fn (rec_prem, ih) => berghofe@20376: let berghofe@21021: val _ $ (S $ x $ (y as Free (_, T))) = berghofe@20376: prop_of rec_prem; berghofe@20376: val k = find_index (equal S) rec_sets; berghofe@21073: val atoms = List.concat (List.mapPartial (fn (bs, z) => berghofe@21073: if z = x then NONE else SOME bs) cargsl') berghofe@20376: in berghofe@21073: map (fn a as Free (_, aT) => berghofe@21073: let val l = find_index (equal aT) dt_atomTs; berghofe@21073: in berghofe@21073: Goal.prove context'' [] [] berghofe@20376: (HOLogic.mk_Trueprop (Const ("Nominal.fresh", berghofe@20376: aT --> T --> HOLogic.boolT) $ a $ y)) berghofe@20376: (fn _ => EVERY berghofe@20376: (rtac (List.nth (List.nth (rec_fresh_thms, l), k)) 1 :: berghofe@20376: map (fn th => rtac th 1) berghofe@20376: (snd (List.nth (finite_thss, l))) @ berghofe@20376: [rtac rec_prem 1, rtac ih 1, berghofe@21073: REPEAT_DETERM (resolve_tac fresh_prems 1)])) berghofe@21073: end) atoms berghofe@21073: end) (rec_prems1 ~~ ihs)); berghofe@20376: berghofe@20376: (** as # fK as ts rs , bs # fK bs us vs **) berghofe@20376: val _ = warning "step 9: as # fK as ts rs , bs # fK bs us vs"; berghofe@21073: fun prove_fresh_result (a as Free (_, aT)) = berghofe@20376: Goal.prove context'' [] [] berghofe@20376: (HOLogic.mk_Trueprop (Const ("Nominal.fresh", berghofe@21073: aT --> rT --> HOLogic.boolT) $ a $ rhs')) berghofe@20376: (fn _ => EVERY berghofe@20376: [resolve_tac fcbs 1, berghofe@20376: REPEAT_DETERM (resolve_tac berghofe@21073: (fresh_prems @ rec_freshs) 1), berghofe@20411: REPEAT_DETERM (resolve_tac (maps snd rec_fin_supp_thms') 1 berghofe@20411: THEN resolve_tac rec_prems 1), berghofe@20376: resolve_tac P_ind_ths 1, berghofe@20376: REPEAT_DETERM (resolve_tac (P_ths @ rec_prems) 1)]); wenzelm@21365: berghofe@21073: val fresh_results'' = map prove_fresh_result boundsl; berghofe@21073: berghofe@21073: fun prove_fresh_result'' ((a as Free (_, aT), b), th) = berghofe@21073: let val th' = Goal.prove context'' [] [] berghofe@21073: (HOLogic.mk_Trueprop (Const ("Nominal.fresh", berghofe@21073: aT --> rT --> HOLogic.boolT) $ berghofe@22311: fold_rev (mk_perm []) (rpi2 @ pi1) a $ berghofe@22311: fold_rev (mk_perm []) (rpi2 @ pi1) rhs')) berghofe@21073: (fn _ => simp_tac (HOL_ss addsimps fresh_bij) 1 THEN berghofe@21073: rtac th 1) berghofe@21073: in berghofe@21073: Goal.prove context'' [] [] berghofe@21073: (HOLogic.mk_Trueprop (Const ("Nominal.fresh", berghofe@21073: aT --> rT --> HOLogic.boolT) $ b $ lhs')) berghofe@21073: (fn _ => EVERY berghofe@21073: [cut_facts_tac [th'] 1, berghofe@21073: NominalPermeq.perm_simp_tac (HOL_ss addsimps berghofe@21073: (rec_eqns @ pi1_pi2_eqs @ perm_swap)) 1, berghofe@21073: full_simp_tac (HOL_ss addsimps (calc_atm @ berghofe@21073: fresh_prems' @ freshs2' @ perm_fresh_fresh)) 1]) berghofe@21073: end; berghofe@21073: berghofe@21073: val fresh_results = fresh_results'' @ map prove_fresh_result'' berghofe@21073: (boundsl ~~ boundsr ~~ fresh_results''); berghofe@20376: berghofe@20376: (** cs # fK as ts rs , cs # fK bs us vs **) berghofe@20376: val _ = warning "step 10: cs # fK as ts rs , cs # fK bs us vs"; berghofe@20376: fun prove_fresh_result' recs t (a as Free (_, aT)) = berghofe@20376: Goal.prove context'' [] [] berghofe@20376: (HOLogic.mk_Trueprop (Const ("Nominal.fresh", berghofe@20376: aT --> rT --> HOLogic.boolT) $ a $ t)) berghofe@20376: (fn _ => EVERY berghofe@20376: [cut_facts_tac recs 1, berghofe@20376: REPEAT_DETERM (dresolve_tac berghofe@20376: (the (AList.lookup op = rec_fin_supp_thms' aT)) 1), berghofe@20376: NominalPermeq.fresh_guess_tac berghofe@20376: (HOL_ss addsimps (freshs2 @ berghofe@20376: fs_atoms @ fresh_atm @ berghofe@20376: List.concat (map snd finite_thss))) 1]); berghofe@20376: berghofe@20376: val fresh_results' = berghofe@20376: map (prove_fresh_result' rec_prems1 rhs') freshs1 @ berghofe@20376: map (prove_fresh_result' rec_prems2 lhs') freshs1; berghofe@20376: berghofe@20376: (** pi1 o (fK as ts rs) = pi2 o (fK bs us vs) **) berghofe@20376: val _ = warning "step 11: pi1 o (fK as ts rs) = pi2 o (fK bs us vs)"; berghofe@20376: val pi1_pi2_result = Goal.prove context'' [] [] berghofe@20376: (HOLogic.mk_Trueprop (HOLogic.mk_eq berghofe@22311: (fold_rev (mk_perm []) pi1 rhs', fold_rev (mk_perm []) pi2 lhs'))) berghofe@20376: (fn _ => NominalPermeq.perm_simp_tac (HOL_ss addsimps berghofe@21073: pi1_pi2_eqs @ rec_eqns) 1 THEN berghofe@20376: TRY (simp_tac (HOL_ss addsimps berghofe@20376: (fresh_prems' @ freshs2' @ calc_atm @ perm_fresh_fresh)) 1)); berghofe@20376: berghofe@20376: val _ = warning "final result"; berghofe@20376: val final = Goal.prove context'' [] [] (term_of concl) berghofe@20376: (fn _ => cut_facts_tac [pi1_pi2_result RS sym] 1 THEN berghofe@20376: full_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh @ berghofe@20376: fresh_results @ fresh_results') 1); berghofe@20376: val final' = ProofContext.export context'' context' [final]; berghofe@20376: val _ = warning "finished!" berghofe@20376: in berghofe@20376: resolve_tac final' 1 berghofe@20397: end) context 1])) idxss) (ndescr ~~ rec_elims)) berghofe@20397: end)); berghofe@20397: berghofe@20397: val rec_total_thms = map (fn r => r RS theI') rec_unique_thms; berghofe@20397: berghofe@20397: (* define primrec combinators *) berghofe@20397: berghofe@20397: val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec"; berghofe@20397: val reccomb_names = map (Sign.full_name thy11) berghofe@20397: (if length descr'' = 1 then [big_reccomb_name] else berghofe@20397: (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int) berghofe@20397: (1 upto (length descr'')))); berghofe@20397: val reccombs = map (fn ((name, T), T') => list_comb berghofe@20397: (Const (name, rec_fn_Ts @ [T] ---> T'), rec_fns)) berghofe@20397: (reccomb_names ~~ recTs ~~ rec_result_Ts); berghofe@20397: berghofe@20397: val (reccomb_defs, thy12) = berghofe@20397: thy11 wenzelm@24712: |> Sign.add_consts_i (map (fn ((name, T), T') => berghofe@20397: (Sign.base_name name, rec_fn_Ts @ [T] ---> T', NoSyn)) berghofe@20397: (reccomb_names ~~ recTs ~~ rec_result_Ts)) berghofe@20397: |> (PureThy.add_defs_i false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') => berghofe@20397: ((Sign.base_name name) ^ "_def", Logic.mk_equals (comb, absfree ("x", T, berghofe@20397: Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T', berghofe@21021: set $ Free ("x", T) $ Free ("y", T')))))) berghofe@20397: (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts)); berghofe@20397: berghofe@20397: (* prove characteristic equations for primrec combinators *) berghofe@20397: berghofe@20397: val rec_thms = map (fn (prems, concl) => berghofe@20397: let berghofe@20397: val _ $ (_ $ (_ $ x) $ _) = concl; berghofe@20397: val (_, cargs) = strip_comb x; berghofe@20397: val ps = map (fn (x as Free (_, T), i) => berghofe@20397: (Free ("x" ^ string_of_int i, T), x)) (cargs ~~ (1 upto length cargs)); berghofe@20397: val concl' = subst_atomic_types (rec_result_Ts' ~~ rec_result_Ts) concl; berghofe@21088: val prems' = List.concat finite_premss @ finite_ctxt_prems @ berghofe@21088: rec_prems @ rec_prems' @ map (subst_atomic ps) prems; berghofe@20397: fun solve rules prems = resolve_tac rules THEN_ALL_NEW berghofe@20397: (resolve_tac prems THEN_ALL_NEW atac) berghofe@20397: in berghofe@20397: Goal.prove_global thy12 [] prems' concl' berghofe@20397: (fn prems => EVERY berghofe@20397: [rewrite_goals_tac reccomb_defs, berghofe@20397: rtac the1_equality 1, berghofe@20397: solve rec_unique_thms prems 1, berghofe@20397: resolve_tac rec_intrs 1, berghofe@20397: REPEAT (solve (prems @ rec_total_thms) prems 1)]) berghofe@20397: end) (rec_eq_prems ~~ berghofe@20397: DatatypeProp.make_primrecs new_type_names descr' sorts' thy12); wenzelm@21365: berghofe@22433: val dt_infos = map (make_dt_info pdescr sorts induct reccomb_names rec_thms) berghofe@21540: ((0 upto length descr1 - 1) ~~ descr1 ~~ distinct_thms ~~ inject_thms); berghofe@21540: berghofe@19985: (* FIXME: theorems are stored in database for testing only *) berghofe@20397: val (_, thy13) = thy12 |> berghofe@20145: PureThy.add_thmss berghofe@20145: [(("rec_equiv", List.concat rec_equiv_thms), []), berghofe@20145: (("rec_equiv'", List.concat rec_equiv_thms'), []), berghofe@20145: (("rec_fin_supp", List.concat rec_fin_supp_thms), []), berghofe@20376: (("rec_fresh", List.concat rec_fresh_thms), []), berghofe@20397: (("rec_unique", map standard rec_unique_thms), []), berghofe@20397: (("recs", rec_thms), [])] ||> wenzelm@24712: Sign.parent_path ||> berghofe@21540: map_nominal_datatypes (fold Symtab.update dt_infos); berghofe@19985: berghofe@17870: in berghofe@20397: thy13 berghofe@17870: end; berghofe@17870: berghofe@17870: val add_nominal_datatype = gen_add_nominal_datatype read_typ true; berghofe@17870: berghofe@17870: berghofe@17870: (* FIXME: The following stuff should be exported by DatatypePackage *) berghofe@17870: berghofe@17870: local structure P = OuterParse and K = OuterKeyword in berghofe@17870: berghofe@17870: val datatype_decl = berghofe@17870: Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix -- berghofe@17870: (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix)); berghofe@17870: berghofe@17870: fun mk_datatype args = berghofe@17870: let berghofe@17870: val names = map (fn ((((NONE, _), t), _), _) => t | ((((SOME t, _), _), _), _) => t) args; berghofe@17870: val specs = map (fn ((((_, vs), t), mx), cons) => berghofe@17870: (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args; berghofe@18068: in add_nominal_datatype false names specs end; berghofe@17870: wenzelm@24867: val _ = berghofe@17870: OuterSyntax.command "nominal_datatype" "define inductive datatypes" K.thy_decl berghofe@17870: (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype)); berghofe@17870: berghofe@17870: end; berghofe@17870: berghofe@18261: end