--- a/NEWS Thu Jun 18 18:31:14 2009 -0700
+++ b/NEWS Fri Jun 19 17:23:21 2009 +0200
@@ -43,6 +43,16 @@
* Constants Set.Pow and Set.image now with authentic syntax; object-logic definitions
Set.Pow_def and Set.image_def. INCOMPATIBILITY.
+* Discontinued ancient tradition to suffix certain ML module names with "_package", e.g.:
+
+ DatatypePackage ~> Datatype
+ InductivePackage ~> Inductive
+
+ etc.
+
+INCOMPATIBILITY.
+
+
*** ML ***
--- a/src/HOL/FunDef.thy Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/FunDef.thy Fri Jun 19 17:23:21 2009 +0200
@@ -17,7 +17,7 @@
("Tools/function_package/sum_tree.ML")
("Tools/function_package/mutual.ML")
("Tools/function_package/pattern_split.ML")
- ("Tools/function_package/fundef_package.ML")
+ ("Tools/function_package/fundef.ML")
("Tools/function_package/auto_term.ML")
("Tools/function_package/measure_functions.ML")
("Tools/function_package/lexicographic_order.ML")
@@ -112,12 +112,12 @@
use "Tools/function_package/mutual.ML"
use "Tools/function_package/pattern_split.ML"
use "Tools/function_package/auto_term.ML"
-use "Tools/function_package/fundef_package.ML"
+use "Tools/function_package/fundef.ML"
use "Tools/function_package/fundef_datatype.ML"
use "Tools/function_package/induction_scheme.ML"
setup {*
- FundefPackage.setup
+ Fundef.setup
#> FundefDatatype.setup
#> InductionScheme.setup
*}
--- a/src/HOL/Hilbert_Choice.thy Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Hilbert_Choice.thy Fri Jun 19 17:23:21 2009 +0200
@@ -7,7 +7,7 @@
theory Hilbert_Choice
imports Nat Wellfounded Plain
-uses ("Tools/meson.ML") ("Tools/specification_package.ML")
+uses ("Tools/meson.ML") ("Tools/choice_specification.ML")
begin
subsection {* Hilbert's epsilon *}
@@ -596,7 +596,7 @@
lemma exE_some: "[| Ex P ; c == Eps P |] ==> P c"
by (simp only: someI_ex)
-use "Tools/specification_package.ML"
+use "Tools/choice_specification.ML"
end
--- a/src/HOL/HoareParallel/OG_Syntax.thy Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/HoareParallel/OG_Syntax.thy Fri Jun 19 17:23:21 2009 +0200
@@ -95,7 +95,7 @@
| annbexp_tr' _ _ = raise Match;
fun upd_tr' (x_upd, T) =
- (case try (unsuffix RecordPackage.updateN) x_upd of
+ (case try (unsuffix Record.updateN) x_upd of
SOME x => (x, if T = dummyT then T else Term.domain_type T)
| NONE => raise Match);
--- a/src/HOL/HoareParallel/RG_Syntax.thy Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/HoareParallel/RG_Syntax.thy Fri Jun 19 17:23:21 2009 +0200
@@ -68,7 +68,7 @@
| bexp_tr' _ _ = raise Match;
fun upd_tr' (x_upd, T) =
- (case try (unsuffix RecordPackage.updateN) x_upd of
+ (case try (unsuffix Record.updateN) x_upd of
SOME x => (x, if T = dummyT then T else Term.domain_type T)
| NONE => raise Match);
--- a/src/HOL/Import/HOL4Setup.thy Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Import/HOL4Setup.thy Fri Jun 19 17:23:21 2009 +0200
@@ -1,10 +1,9 @@
(* Title: HOL/Import/HOL4Setup.thy
- ID: $Id$
Author: Sebastian Skalberg (TU Muenchen)
*)
theory HOL4Setup imports MakeEqual ImportRecorder
- uses ("proof_kernel.ML") ("replay.ML") ("hol4rews.ML") ("import_package.ML") begin
+ uses ("proof_kernel.ML") ("replay.ML") ("hol4rews.ML") ("import.ML") begin
section {* General Setup *}
@@ -162,8 +161,8 @@
use "proof_kernel.ML"
use "replay.ML"
-use "import_package.ML"
+use "import.ML"
-setup ImportPackage.setup
+setup Import.setup
end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Import/import.ML Fri Jun 19 17:23:21 2009 +0200
@@ -0,0 +1,71 @@
+(* Title: HOL/Import/import.ML
+ Author: Sebastian Skalberg (TU Muenchen)
+*)
+
+signature IMPORT =
+sig
+ val debug : bool ref
+ val import_tac : Proof.context -> string * string -> tactic
+ val setup : theory -> theory
+end
+
+structure ImportData = TheoryDataFun
+(
+ type T = ProofKernel.thm option array option
+ val empty = NONE
+ val copy = I
+ val extend = I
+ fun merge _ _ = NONE
+)
+
+structure Import :> IMPORT =
+struct
+
+val debug = ref false
+fun message s = if !debug then writeln s else ()
+
+fun import_tac ctxt (thyname, thmname) =
+ if ! quick_and_dirty
+ then SkipProof.cheat_tac (ProofContext.theory_of ctxt)
+ else
+ fn th =>
+ let
+ val thy = ProofContext.theory_of ctxt
+ val prem = hd (prems_of th)
+ val _ = message ("Import_tac: thyname=" ^ thyname ^ ", thmname=" ^ thmname)
+ val _ = message ("Import trying to prove " ^ Syntax.string_of_term ctxt prem)
+ val int_thms = case ImportData.get thy of
+ NONE => fst (Replay.setup_int_thms thyname thy)
+ | SOME a => a
+ val proof = snd (ProofKernel.import_proof thyname thmname thy) thy
+ val hol4thm = snd (Replay.replay_proof int_thms thyname thmname proof thy)
+ val thm = snd (ProofKernel.to_isa_thm hol4thm)
+ val rew = ProofKernel.rewrite_hol4_term (concl_of thm) thy
+ val thm = equal_elim rew thm
+ val prew = ProofKernel.rewrite_hol4_term prem thy
+ val prem' = #2 (Logic.dest_equals (prop_of prew))
+ val _ = message ("Import proved " ^ Display.string_of_thm thm)
+ val thm = ProofKernel.disambiguate_frees thm
+ val _ = message ("Disambiguate: " ^ Display.string_of_thm thm)
+ in
+ case Shuffler.set_prop thy prem' [("",thm)] of
+ SOME (_,thm) =>
+ let
+ val _ = if prem' aconv (prop_of thm)
+ then message "import: Terms match up"
+ else message "import: Terms DO NOT match up"
+ val thm' = equal_elim (symmetric prew) thm
+ val res = bicompose true (false,thm',0) 1 th
+ in
+ res
+ end
+ | NONE => (message "import: set_prop didn't succeed"; no_tac th)
+ end
+
+val setup = Method.setup @{binding import}
+ (Scan.lift (Args.name -- Args.name) >>
+ (fn arg => fn ctxt => SIMPLE_METHOD (import_tac ctxt arg)))
+ "import HOL4 theorem"
+
+end
+
--- a/src/HOL/Import/import_package.ML Thu Jun 18 18:31:14 2009 -0700
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,71 +0,0 @@
-(* Title: HOL/Import/import_package.ML
- Author: Sebastian Skalberg (TU Muenchen)
-*)
-
-signature IMPORT_PACKAGE =
-sig
- val debug : bool ref
- val import_tac : Proof.context -> string * string -> tactic
- val setup : theory -> theory
-end
-
-structure ImportData = TheoryDataFun
-(
- type T = ProofKernel.thm option array option
- val empty = NONE
- val copy = I
- val extend = I
- fun merge _ _ = NONE
-)
-
-structure ImportPackage :> IMPORT_PACKAGE =
-struct
-
-val debug = ref false
-fun message s = if !debug then writeln s else ()
-
-fun import_tac ctxt (thyname, thmname) =
- if ! quick_and_dirty
- then SkipProof.cheat_tac (ProofContext.theory_of ctxt)
- else
- fn th =>
- let
- val thy = ProofContext.theory_of ctxt
- val prem = hd (prems_of th)
- val _ = message ("Import_tac: thyname=" ^ thyname ^ ", thmname=" ^ thmname)
- val _ = message ("Import trying to prove " ^ Syntax.string_of_term ctxt prem)
- val int_thms = case ImportData.get thy of
- NONE => fst (Replay.setup_int_thms thyname thy)
- | SOME a => a
- val proof = snd (ProofKernel.import_proof thyname thmname thy) thy
- val hol4thm = snd (Replay.replay_proof int_thms thyname thmname proof thy)
- val thm = snd (ProofKernel.to_isa_thm hol4thm)
- val rew = ProofKernel.rewrite_hol4_term (concl_of thm) thy
- val thm = equal_elim rew thm
- val prew = ProofKernel.rewrite_hol4_term prem thy
- val prem' = #2 (Logic.dest_equals (prop_of prew))
- val _ = message ("Import proved " ^ Display.string_of_thm thm)
- val thm = ProofKernel.disambiguate_frees thm
- val _ = message ("Disambiguate: " ^ Display.string_of_thm thm)
- in
- case Shuffler.set_prop thy prem' [("",thm)] of
- SOME (_,thm) =>
- let
- val _ = if prem' aconv (prop_of thm)
- then message "import: Terms match up"
- else message "import: Terms DO NOT match up"
- val thm' = equal_elim (symmetric prew) thm
- val res = bicompose true (false,thm',0) 1 th
- in
- res
- end
- | NONE => (message "import: set_prop didn't succeed"; no_tac th)
- end
-
-val setup = Method.setup @{binding import}
- (Scan.lift (Args.name -- Args.name) >>
- (fn arg => fn ctxt => SIMPLE_METHOD (import_tac ctxt arg)))
- "import HOL4 theorem"
-
-end
-
--- a/src/HOL/Import/proof_kernel.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Import/proof_kernel.ML Fri Jun 19 17:23:21 2009 +0200
@@ -2021,7 +2021,7 @@
snd (get_defname thyname name thy)) thy1 names
fun new_name name = fst (get_defname thyname name thy1)
val names' = map (fn name => (new_name name,name,false)) names
- val (thy',res) = SpecificationPackage.add_specification NONE
+ val (thy',res) = Choice_Specification.add_specification NONE
names'
(thy1,th)
val _ = ImportRecorder.add_specification names' th
@@ -2091,7 +2091,7 @@
val tsyn = mk_syn thy tycname
val typ = (tycname,tnames,tsyn)
val ((_, typedef_info), thy') =
- TypedefPackage.add_typedef false (SOME (Binding.name thmname))
+ Typedef.add_typedef false (SOME (Binding.name thmname))
(Binding.name tycname, tnames, tsyn) c NONE (rtac th2 1) thy
val _ = ImportRecorder.add_typedef (SOME thmname) typ c NONE th2
@@ -2179,7 +2179,7 @@
val tsyn = mk_syn thy tycname
val typ = (tycname,tnames,tsyn)
val ((_, typedef_info), thy') =
- TypedefPackage.add_typedef false NONE (Binding.name tycname,tnames,tsyn) c
+ Typedef.add_typedef false NONE (Binding.name tycname,tnames,tsyn) c
(SOME(Binding.name rep_name,Binding.name abs_name)) (rtac th2 1) thy
val _ = ImportRecorder.add_typedef NONE typ c (SOME(rep_name,abs_name)) th2
val fulltyname = Sign.intern_type thy' tycname
--- a/src/HOL/Import/replay.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Import/replay.ML Fri Jun 19 17:23:21 2009 +0200
@@ -329,7 +329,7 @@
and rp (ThmEntry (thyname', thmname', aborted, History history)) thy = rps history thy
| rp (DeltaEntry ds) thy = fold delta ds thy
and delta (Specification (names, th)) thy =
- fst (SpecificationPackage.add_specification NONE names (thy,th_of thy th))
+ fst (Choice_Specification.add_specification NONE names (thy,th_of thy th))
| delta (Hol_mapping (thyname, thmname, isaname)) thy =
add_hol4_mapping thyname thmname isaname thy
| delta (Hol_pending (thyname, thmname, th)) thy =
@@ -344,7 +344,7 @@
| delta (Hol_theorem (thyname, thmname, th)) thy =
add_hol4_theorem thyname thmname ([], th_of thy th) thy
| delta (Typedef (thmname, (t, vs, mx), c, repabs, th)) thy =
- snd (TypedefPackage.add_typedef false (Option.map Binding.name thmname) (Binding.name t, vs, mx) c
+ snd (Typedef.add_typedef false (Option.map Binding.name thmname) (Binding.name t, vs, mx) c
(Option.map (pairself Binding.name) repabs) (rtac (th_of thy th) 1) thy)
| delta (Hol_type_mapping (thyname, tycname, fulltyname)) thy =
add_hol4_type_mapping thyname tycname true fulltyname thy
--- a/src/HOL/Inductive.thy Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Inductive.thy Fri Jun 19 17:23:21 2009 +0200
@@ -7,7 +7,7 @@
theory Inductive
imports Lattices Sum_Type
uses
- ("Tools/inductive_package.ML")
+ ("Tools/inductive.ML")
"Tools/dseq.ML"
("Tools/inductive_codegen.ML")
("Tools/datatype_package/datatype_aux.ML")
@@ -15,9 +15,9 @@
("Tools/datatype_package/datatype_rep_proofs.ML")
("Tools/datatype_package/datatype_abs_proofs.ML")
("Tools/datatype_package/datatype_case.ML")
- ("Tools/datatype_package/datatype_package.ML")
- ("Tools/old_primrec_package.ML")
- ("Tools/primrec_package.ML")
+ ("Tools/datatype_package/datatype.ML")
+ ("Tools/old_primrec.ML")
+ ("Tools/primrec.ML")
("Tools/datatype_package/datatype_codegen.ML")
begin
@@ -320,8 +320,8 @@
val le_fun_def = @{thm le_fun_def} RS @{thm eq_reflection}
*}
-use "Tools/inductive_package.ML"
-setup InductivePackage.setup
+use "Tools/inductive.ML"
+setup Inductive.setup
theorems [mono] =
imp_refl disj_mono conj_mono ex_mono all_mono if_bool_eq_conj
@@ -340,11 +340,11 @@
use "Tools/datatype_package/datatype_rep_proofs.ML"
use "Tools/datatype_package/datatype_abs_proofs.ML"
use "Tools/datatype_package/datatype_case.ML"
-use "Tools/datatype_package/datatype_package.ML"
-setup DatatypePackage.setup
+use "Tools/datatype_package/datatype.ML"
+setup Datatype.setup
-use "Tools/old_primrec_package.ML"
-use "Tools/primrec_package.ML"
+use "Tools/old_primrec.ML"
+use "Tools/primrec.ML"
use "Tools/datatype_package/datatype_codegen.ML"
setup DatatypeCodegen.setup
@@ -364,7 +364,7 @@
fun fun_tr ctxt [cs] =
let
val x = Free (Name.variant (Term.add_free_names cs []) "x", dummyT);
- val ft = DatatypeCase.case_tr true DatatypePackage.datatype_of_constr
+ val ft = DatatypeCase.case_tr true Datatype.datatype_of_constr
ctxt [x, cs]
in lambda x ft end
in [("_lam_pats_syntax", fun_tr)] end
--- a/src/HOL/IsaMakefile Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/IsaMakefile Fri Jun 19 17:23:21 2009 +0200
@@ -145,7 +145,7 @@
Tools/datatype_package/datatype_aux.ML \
Tools/datatype_package/datatype_case.ML \
Tools/datatype_package/datatype_codegen.ML \
- Tools/datatype_package/datatype_package.ML \
+ Tools/datatype_package/datatype.ML \
Tools/datatype_package/datatype_prop.ML \
Tools/datatype_package/datatype_realizer.ML \
Tools/datatype_package/datatype_rep_proofs.ML \
@@ -158,7 +158,7 @@
Tools/function_package/fundef_core.ML \
Tools/function_package/fundef_datatype.ML \
Tools/function_package/fundef_lib.ML \
- Tools/function_package/fundef_package.ML \
+ Tools/function_package/fundef.ML \
Tools/function_package/induction_scheme.ML \
Tools/function_package/inductive_wrap.ML \
Tools/function_package/lexicographic_order.ML \
@@ -171,24 +171,24 @@
Tools/function_package/sum_tree.ML \
Tools/function_package/termination.ML \
Tools/inductive_codegen.ML \
- Tools/inductive_package.ML \
+ Tools/inductive.ML \
Tools/inductive_realizer.ML \
- Tools/inductive_set_package.ML \
+ Tools/inductive_set.ML \
Tools/lin_arith.ML \
Tools/nat_arith.ML \
- Tools/old_primrec_package.ML \
- Tools/primrec_package.ML \
+ Tools/old_primrec.ML \
+ Tools/primrec.ML \
Tools/prop_logic.ML \
- Tools/record_package.ML \
+ Tools/record.ML \
Tools/refute.ML \
Tools/refute_isar.ML \
Tools/rewrite_hol_proof.ML \
Tools/sat_funcs.ML \
Tools/sat_solver.ML \
Tools/split_rule.ML \
- Tools/typecopy_package.ML \
+ Tools/typecopy.ML \
Tools/typedef_codegen.ML \
- Tools/typedef_package.ML \
+ Tools/typedef.ML \
Transitive_Closure.thy \
Typedef.thy \
Wellfounded.thy \
@@ -250,13 +250,13 @@
Tools/Qelim/generated_cooper.ML \
Tools/Qelim/presburger.ML \
Tools/Qelim/qelim.ML \
- Tools/recdef_package.ML \
+ Tools/recdef.ML \
Tools/res_atp.ML \
Tools/res_axioms.ML \
Tools/res_clause.ML \
Tools/res_hol_clause.ML \
Tools/res_reconstruct.ML \
- Tools/specification_package.ML \
+ Tools/choice_specification.ML \
Tools/string_code.ML \
Tools/string_syntax.ML \
Tools/TFL/casesplit.ML \
@@ -423,7 +423,7 @@
IMPORTER_FILES = Import/lazy_seq.ML Import/proof_kernel.ML Import/replay.ML \
Import/shuffler.ML Import/MakeEqual.thy Import/HOL4Setup.thy \
Import/HOL4Syntax.thy Import/HOL4Compat.thy Import/import_syntax.ML \
- Import/hol4rews.ML Import/import_package.ML Import/ROOT.ML
+ Import/hol4rews.ML Import/import.ML Import/ROOT.ML
IMPORTER_HOLLIGHT_FILES = Import/proof_kernel.ML Import/replay.ML \
Import/shuffler.ML Import/MakeEqual.thy Import/HOL4Setup.thy \
@@ -968,7 +968,7 @@
Nominal/nominal_induct.ML \
Nominal/nominal_inductive.ML \
Nominal/nominal_inductive2.ML \
- Nominal/nominal_package.ML \
+ Nominal/nominal.ML \
Nominal/nominal_permeq.ML \
Nominal/nominal_primrec.ML \
Nominal/nominal_thmdecls.ML \
--- a/src/HOL/Isar_examples/Hoare.thy Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Isar_examples/Hoare.thy Fri Jun 19 17:23:21 2009 +0200
@@ -260,7 +260,7 @@
| bexp_tr' _ _ = raise Match;
fun upd_tr' (x_upd, T) =
- (case try (unsuffix RecordPackage.updateN) x_upd of
+ (case try (unsuffix Record.updateN) x_upd of
SOME x => (x, if T = dummyT then T else Term.domain_type T)
| NONE => raise Match);
--- a/src/HOL/List.thy Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/List.thy Fri Jun 19 17:23:21 2009 +0200
@@ -363,7 +363,7 @@
val case2 = Syntax.const "_case1" $ Syntax.const Term.dummy_patternN
$ NilC;
val cs = Syntax.const "_case2" $ case1 $ case2
- val ft = DatatypeCase.case_tr false DatatypePackage.datatype_of_constr
+ val ft = DatatypeCase.case_tr false Datatype.datatype_of_constr
ctxt [x, cs]
in lambda x ft end;
--- a/src/HOL/Nominal/Nominal.thy Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Nominal/Nominal.thy Fri Jun 19 17:23:21 2009 +0200
@@ -3,7 +3,7 @@
uses
("nominal_thmdecls.ML")
("nominal_atoms.ML")
- ("nominal_package.ML")
+ ("nominal.ML")
("nominal_induct.ML")
("nominal_permeq.ML")
("nominal_fresh_fun.ML")
@@ -3670,7 +3670,7 @@
lemma allE_Nil: assumes "\<forall>x. P x" obtains "P []"
using assms ..
-use "nominal_package.ML"
+use "nominal.ML"
(******************************************************)
(* primitive recursive functions on nominal datatypes *)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Nominal/nominal.ML Fri Jun 19 17:23:21 2009 +0200
@@ -0,0 +1,2095 @@
+(* Title: HOL/Nominal/nominal.ML
+ Author: Stefan Berghofer and Christian Urban, TU Muenchen
+
+Nominal datatype package for Isabelle/HOL.
+*)
+
+signature NOMINAL =
+sig
+ val add_nominal_datatype : DatatypeAux.datatype_config -> string list ->
+ (string list * bstring * mixfix *
+ (bstring * string list * mixfix) list) list -> theory -> theory
+ type descr
+ type nominal_datatype_info
+ val get_nominal_datatypes : theory -> nominal_datatype_info Symtab.table
+ val get_nominal_datatype : theory -> string -> nominal_datatype_info option
+ val mk_perm: typ list -> term -> term -> term
+ val perm_of_pair: term * term -> term
+ val mk_not_sym: thm list -> thm list
+ val perm_simproc: simproc
+ val fresh_const: typ -> typ -> term
+ val fresh_star_const: typ -> typ -> term
+end
+
+structure Nominal : NOMINAL =
+struct
+
+val finite_emptyI = thm "finite.emptyI";
+val finite_Diff = thm "finite_Diff";
+val finite_Un = thm "finite_Un";
+val Un_iff = thm "Un_iff";
+val In0_eq = thm "In0_eq";
+val In1_eq = thm "In1_eq";
+val In0_not_In1 = thm "In0_not_In1";
+val In1_not_In0 = thm "In1_not_In0";
+val Un_assoc = thm "Un_assoc";
+val Collect_disj_eq = thm "Collect_disj_eq";
+val empty_def = thm "empty_def";
+val empty_iff = thm "empty_iff";
+
+open DatatypeAux;
+open NominalAtoms;
+
+(** FIXME: Datatype should export this function **)
+
+local
+
+fun dt_recs (DtTFree _) = []
+ | dt_recs (DtType (_, dts)) = List.concat (map dt_recs dts)
+ | dt_recs (DtRec i) = [i];
+
+fun dt_cases (descr: descr) (_, args, constrs) =
+ let
+ fun the_bname i = Long_Name.base_name (#1 (valOf (AList.lookup (op =) descr i)));
+ val bnames = map the_bname (distinct op = (List.concat (map dt_recs args)));
+ in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end;
+
+
+fun induct_cases descr =
+ DatatypeProp.indexify_names (List.concat (map (dt_cases descr) (map #2 descr)));
+
+fun exhaust_cases descr i = dt_cases descr (valOf (AList.lookup (op =) descr i));
+
+in
+
+fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
+
+fun mk_case_names_exhausts descr new =
+ map (RuleCases.case_names o exhaust_cases descr o #1)
+ (List.filter (fn ((_, (name, _, _))) => name mem_string new) descr);
+
+end;
+
+(* theory data *)
+
+type descr = (int * (string * dtyp list * (string * (dtyp list * dtyp) list) list)) list;
+
+type nominal_datatype_info =
+ {index : int,
+ descr : descr,
+ sorts : (string * sort) list,
+ rec_names : string list,
+ rec_rewrites : thm list,
+ induction : thm,
+ distinct : thm list,
+ inject : thm list};
+
+structure NominalDatatypesData = TheoryDataFun
+(
+ type T = nominal_datatype_info Symtab.table;
+ val empty = Symtab.empty;
+ val copy = I;
+ val extend = I;
+ fun merge _ tabs : T = Symtab.merge (K true) tabs;
+);
+
+val get_nominal_datatypes = NominalDatatypesData.get;
+val put_nominal_datatypes = NominalDatatypesData.put;
+val map_nominal_datatypes = NominalDatatypesData.map;
+val get_nominal_datatype = Symtab.lookup o get_nominal_datatypes;
+
+
+(**** make datatype info ****)
+
+fun make_dt_info descr sorts induct reccomb_names rec_thms
+ (((i, (_, (tname, _, _))), distinct), inject) =
+ (tname,
+ {index = i,
+ descr = descr,
+ sorts = sorts,
+ rec_names = reccomb_names,
+ rec_rewrites = rec_thms,
+ induction = induct,
+ distinct = distinct,
+ inject = inject});
+
+(*******************************)
+
+val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
+
+
+(** simplification procedure for sorting permutations **)
+
+val dj_cp = thm "dj_cp";
+
+fun dest_permT (Type ("fun", [Type ("List.list", [Type ("*", [T, _])]),
+ Type ("fun", [_, U])])) = (T, U);
+
+fun permTs_of (Const ("Nominal.perm", T) $ t $ u) = fst (dest_permT T) :: permTs_of u
+ | permTs_of _ = [];
+
+fun perm_simproc' thy ss (Const ("Nominal.perm", T) $ t $ (u as Const ("Nominal.perm", U) $ r $ s)) =
+ let
+ val (aT as Type (a, []), S) = dest_permT T;
+ val (bT as Type (b, []), _) = dest_permT U
+ in if aT mem permTs_of u andalso aT <> bT then
+ let
+ val cp = cp_inst_of thy a b;
+ val dj = dj_thm_of thy b a;
+ val dj_cp' = [cp, dj] MRS dj_cp;
+ val cert = SOME o cterm_of thy
+ in
+ SOME (mk_meta_eq (Drule.instantiate' [SOME (ctyp_of thy S)]
+ [cert t, cert r, cert s] dj_cp'))
+ end
+ else NONE
+ end
+ | perm_simproc' thy ss _ = NONE;
+
+val perm_simproc =
+ Simplifier.simproc (the_context ()) "perm_simp" ["pi1 \<bullet> (pi2 \<bullet> x)"] perm_simproc';
+
+val meta_spec = thm "meta_spec";
+
+fun projections rule =
+ ProjectRule.projections (ProofContext.init (Thm.theory_of_thm rule)) rule
+ |> map (standard #> RuleCases.save rule);
+
+val supp_prod = thm "supp_prod";
+val fresh_prod = thm "fresh_prod";
+val supports_fresh = thm "supports_fresh";
+val supports_def = thm "Nominal.supports_def";
+val fresh_def = thm "fresh_def";
+val supp_def = thm "supp_def";
+val rev_simps = thms "rev.simps";
+val app_simps = thms "append.simps";
+val at_fin_set_supp = thm "at_fin_set_supp";
+val at_fin_set_fresh = thm "at_fin_set_fresh";
+val abs_fun_eq1 = thm "abs_fun_eq1";
+
+val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq];
+
+fun mk_perm Ts t u =
+ let
+ val T = fastype_of1 (Ts, t);
+ val U = fastype_of1 (Ts, u)
+ in Const ("Nominal.perm", T --> U --> U) $ t $ u end;
+
+fun perm_of_pair (x, y) =
+ let
+ val T = fastype_of x;
+ val pT = mk_permT T
+ in Const ("List.list.Cons", HOLogic.mk_prodT (T, T) --> pT --> pT) $
+ HOLogic.mk_prod (x, y) $ Const ("List.list.Nil", pT)
+ end;
+
+fun mk_not_sym ths = maps (fn th => case prop_of th of
+ _ $ (Const ("Not", _) $ (Const ("op =", _) $ _ $ _)) => [th, th RS not_sym]
+ | _ => [th]) ths;
+
+fun fresh_const T U = Const ("Nominal.fresh", T --> U --> HOLogic.boolT);
+fun fresh_star_const T U =
+ Const ("Nominal.fresh_star", HOLogic.mk_setT T --> U --> HOLogic.boolT);
+
+fun gen_add_nominal_datatype prep_typ config new_type_names dts thy =
+ let
+ (* this theory is used just for parsing *)
+
+ val tmp_thy = thy |>
+ Theory.copy |>
+ Sign.add_types (map (fn (tvs, tname, mx, _) =>
+ (Binding.name tname, length tvs, mx)) dts);
+
+ val atoms = atoms_of thy;
+
+ fun prep_constr ((constrs, sorts), (cname, cargs, mx)) =
+ let val (cargs', sorts') = Library.foldl (prep_typ tmp_thy) (([], sorts), cargs)
+ in (constrs @ [(cname, cargs', mx)], sorts') end
+
+ fun prep_dt_spec ((dts, sorts), (tvs, tname, mx, constrs)) =
+ let val (constrs', sorts') = Library.foldl prep_constr (([], sorts), constrs)
+ in (dts @ [(tvs, tname, mx, constrs')], sorts') end
+
+ val (dts', sorts) = Library.foldl prep_dt_spec (([], []), dts);
+ val tyvars = map (map (fn s =>
+ (s, the (AList.lookup (op =) sorts s))) o #1) dts';
+
+ fun inter_sort thy S S' = Type.inter_sort (Sign.tsig_of thy) (S, S');
+ fun augment_sort_typ thy S =
+ let val S = Sign.certify_sort thy S
+ in map_type_tfree (fn (s, S') => TFree (s,
+ if member (op = o apsnd fst) sorts s then inter_sort thy S S' else S'))
+ end;
+ fun augment_sort thy S = map_types (augment_sort_typ thy S);
+
+ val types_syntax = map (fn (tvs, tname, mx, constrs) => (tname, mx)) dts';
+ val constr_syntax = map (fn (tvs, tname, mx, constrs) =>
+ map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts';
+
+ val ps = map (fn (_, n, _, _) =>
+ (Sign.full_bname tmp_thy n, Sign.full_bname tmp_thy (n ^ "_Rep"))) dts;
+ val rps = map Library.swap ps;
+
+ fun replace_types (Type ("Nominal.ABS", [T, U])) =
+ Type ("fun", [T, Type ("Nominal.noption", [replace_types U])])
+ | replace_types (Type (s, Ts)) =
+ Type (getOpt (AList.lookup op = ps s, s), map replace_types Ts)
+ | replace_types T = T;
+
+ val dts'' = map (fn (tvs, tname, mx, constrs) => (tvs, Binding.name (tname ^ "_Rep"), NoSyn,
+ map (fn (cname, cargs, mx) => (Binding.name (cname ^ "_Rep"),
+ map replace_types cargs, NoSyn)) constrs)) dts';
+
+ val new_type_names' = map (fn n => n ^ "_Rep") new_type_names;
+ val full_new_type_names' = map (Sign.full_bname thy) new_type_names';
+
+ val ({induction, ...},thy1) =
+ Datatype.add_datatype config new_type_names' dts'' thy;
+
+ val SOME {descr, ...} = Symtab.lookup
+ (Datatype.get_datatypes thy1) (hd full_new_type_names');
+ fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
+
+ val big_name = space_implode "_" new_type_names;
+
+
+ (**** define permutation functions ****)
+
+ val permT = mk_permT (TFree ("'x", HOLogic.typeS));
+ val pi = Free ("pi", permT);
+ val perm_types = map (fn (i, _) =>
+ let val T = nth_dtyp i
+ in permT --> T --> T end) descr;
+ val perm_names' = DatatypeProp.indexify_names (map (fn (i, _) =>
+ "perm_" ^ name_of_typ (nth_dtyp i)) descr);
+ val perm_names = replicate (length new_type_names) "Nominal.perm" @
+ map (Sign.full_bname thy1) (List.drop (perm_names', length new_type_names));
+ val perm_names_types = perm_names ~~ perm_types;
+ val perm_names_types' = perm_names' ~~ perm_types;
+
+ val perm_eqs = maps (fn (i, (_, _, constrs)) =>
+ let val T = nth_dtyp i
+ in map (fn (cname, dts) =>
+ let
+ val Ts = map (typ_of_dtyp descr sorts) dts;
+ val names = Name.variant_list ["pi"] (DatatypeProp.make_tnames Ts);
+ val args = map Free (names ~~ Ts);
+ val c = Const (cname, Ts ---> T);
+ fun perm_arg (dt, x) =
+ let val T = type_of x
+ in if is_rec_type dt then
+ let val (Us, _) = strip_type T
+ in list_abs (map (pair "x") Us,
+ Free (nth perm_names_types' (body_index dt)) $ pi $
+ list_comb (x, map (fn (i, U) =>
+ Const ("Nominal.perm", permT --> U --> U) $
+ (Const ("List.rev", permT --> permT) $ pi) $
+ Bound i) ((length Us - 1 downto 0) ~~ Us)))
+ end
+ else Const ("Nominal.perm", permT --> T --> T) $ pi $ x
+ end;
+ in
+ (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (Free (nth perm_names_types' i) $
+ Free ("pi", mk_permT (TFree ("'x", HOLogic.typeS))) $
+ list_comb (c, args),
+ list_comb (c, map perm_arg (dts ~~ args)))))
+ end) constrs
+ end) descr;
+
+ val (perm_simps, thy2) =
+ Primrec.add_primrec_overloaded
+ (map (fn (s, sT) => (s, sT, false))
+ (List.take (perm_names' ~~ perm_names_types, length new_type_names)))
+ (map (fn s => (Binding.name s, NONE, NoSyn)) perm_names') perm_eqs thy1;
+
+ (**** prove that permutation functions introduced by unfolding are ****)
+ (**** equivalent to already existing permutation functions ****)
+
+ val _ = warning ("length descr: " ^ string_of_int (length descr));
+ val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names));
+
+ val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types);
+ val perm_fun_def = PureThy.get_thm thy2 "perm_fun_def";
+
+ val unfolded_perm_eq_thms =
+ if length descr = length new_type_names then []
+ else map standard (List.drop (split_conj_thm
+ (Goal.prove_global thy2 [] []
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+ (map (fn (c as (s, T), x) =>
+ let val [T1, T2] = binder_types T
+ in HOLogic.mk_eq (Const c $ pi $ Free (x, T2),
+ Const ("Nominal.perm", T) $ pi $ Free (x, T2))
+ end)
+ (perm_names_types ~~ perm_indnames))))
+ (fn _ => EVERY [indtac induction perm_indnames 1,
+ ALLGOALS (asm_full_simp_tac
+ (simpset_of thy2 addsimps [perm_fun_def]))])),
+ length new_type_names));
+
+ (**** prove [] \<bullet> t = t ****)
+
+ val _ = warning "perm_empty_thms";
+
+ val perm_empty_thms = List.concat (map (fn a =>
+ let val permT = mk_permT (Type (a, []))
+ in map standard (List.take (split_conj_thm
+ (Goal.prove_global thy2 [] []
+ (augment_sort thy2 [pt_class_of thy2 a]
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+ (map (fn ((s, T), x) => HOLogic.mk_eq
+ (Const (s, permT --> T --> T) $
+ Const ("List.list.Nil", permT) $ Free (x, T),
+ Free (x, T)))
+ (perm_names ~~
+ map body_type perm_types ~~ perm_indnames)))))
+ (fn _ => EVERY [indtac induction perm_indnames 1,
+ ALLGOALS (asm_full_simp_tac (simpset_of thy2))])),
+ length new_type_names))
+ end)
+ atoms);
+
+ (**** prove (pi1 @ pi2) \<bullet> t = pi1 \<bullet> (pi2 \<bullet> t) ****)
+
+ val _ = warning "perm_append_thms";
+
+ (*FIXME: these should be looked up statically*)
+ val at_pt_inst = PureThy.get_thm thy2 "at_pt_inst";
+ val pt2 = PureThy.get_thm thy2 "pt2";
+
+ val perm_append_thms = List.concat (map (fn a =>
+ let
+ val permT = mk_permT (Type (a, []));
+ val pi1 = Free ("pi1", permT);
+ val pi2 = Free ("pi2", permT);
+ val pt_inst = pt_inst_of thy2 a;
+ val pt2' = pt_inst RS pt2;
+ val pt2_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "2") a);
+ in List.take (map standard (split_conj_thm
+ (Goal.prove_global thy2 [] []
+ (augment_sort thy2 [pt_class_of thy2 a]
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+ (map (fn ((s, T), x) =>
+ let val perm = Const (s, permT --> T --> T)
+ in HOLogic.mk_eq
+ (perm $ (Const ("List.append", permT --> permT --> permT) $
+ pi1 $ pi2) $ Free (x, T),
+ perm $ pi1 $ (perm $ pi2 $ Free (x, T)))
+ end)
+ (perm_names ~~
+ map body_type perm_types ~~ perm_indnames)))))
+ (fn _ => EVERY [indtac induction perm_indnames 1,
+ ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt2', pt2_ax]))]))),
+ length new_type_names)
+ end) atoms);
+
+ (**** prove pi1 ~ pi2 ==> pi1 \<bullet> t = pi2 \<bullet> t ****)
+
+ val _ = warning "perm_eq_thms";
+
+ val pt3 = PureThy.get_thm thy2 "pt3";
+ val pt3_rev = PureThy.get_thm thy2 "pt3_rev";
+
+ val perm_eq_thms = List.concat (map (fn a =>
+ let
+ val permT = mk_permT (Type (a, []));
+ val pi1 = Free ("pi1", permT);
+ val pi2 = Free ("pi2", permT);
+ val at_inst = at_inst_of thy2 a;
+ val pt_inst = pt_inst_of thy2 a;
+ val pt3' = pt_inst RS pt3;
+ val pt3_rev' = at_inst RS (pt_inst RS pt3_rev);
+ val pt3_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "3") a);
+ in List.take (map standard (split_conj_thm
+ (Goal.prove_global thy2 [] []
+ (augment_sort thy2 [pt_class_of thy2 a] (Logic.mk_implies
+ (HOLogic.mk_Trueprop (Const ("Nominal.prm_eq",
+ permT --> permT --> HOLogic.boolT) $ pi1 $ pi2),
+ HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+ (map (fn ((s, T), x) =>
+ let val perm = Const (s, permT --> T --> T)
+ in HOLogic.mk_eq
+ (perm $ pi1 $ Free (x, T),
+ perm $ pi2 $ Free (x, T))
+ end)
+ (perm_names ~~
+ map body_type perm_types ~~ perm_indnames))))))
+ (fn _ => EVERY [indtac induction perm_indnames 1,
+ ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))),
+ length new_type_names)
+ end) atoms);
+
+ (**** prove pi1 \<bullet> (pi2 \<bullet> t) = (pi1 \<bullet> pi2) \<bullet> (pi1 \<bullet> t) ****)
+
+ val cp1 = PureThy.get_thm thy2 "cp1";
+ val dj_cp = PureThy.get_thm thy2 "dj_cp";
+ val pt_perm_compose = PureThy.get_thm thy2 "pt_perm_compose";
+ val pt_perm_compose_rev = PureThy.get_thm thy2 "pt_perm_compose_rev";
+ val dj_perm_perm_forget = PureThy.get_thm thy2 "dj_perm_perm_forget";
+
+ fun composition_instance name1 name2 thy =
+ let
+ val cp_class = cp_class_of thy name1 name2;
+ val pt_class =
+ if name1 = name2 then [pt_class_of thy name1]
+ else [];
+ val permT1 = mk_permT (Type (name1, []));
+ val permT2 = mk_permT (Type (name2, []));
+ val Ts = map body_type perm_types;
+ val cp_inst = cp_inst_of thy name1 name2;
+ val simps = simpset_of thy addsimps (perm_fun_def ::
+ (if name1 <> name2 then
+ let val dj = dj_thm_of thy name2 name1
+ in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end
+ else
+ let
+ val at_inst = at_inst_of thy name1;
+ val pt_inst = pt_inst_of thy name1;
+ in
+ [cp_inst RS cp1 RS sym,
+ at_inst RS (pt_inst RS pt_perm_compose) RS sym,
+ at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym]
+ end))
+ val sort = Sign.certify_sort thy (cp_class :: pt_class);
+ val thms = split_conj_thm (Goal.prove_global thy [] []
+ (augment_sort thy sort
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+ (map (fn ((s, T), x) =>
+ let
+ val pi1 = Free ("pi1", permT1);
+ val pi2 = Free ("pi2", permT2);
+ val perm1 = Const (s, permT1 --> T --> T);
+ val perm2 = Const (s, permT2 --> T --> T);
+ val perm3 = Const ("Nominal.perm", permT1 --> permT2 --> permT2)
+ in HOLogic.mk_eq
+ (perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)),
+ perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T)))
+ end)
+ (perm_names ~~ Ts ~~ perm_indnames)))))
+ (fn _ => EVERY [indtac induction perm_indnames 1,
+ ALLGOALS (asm_full_simp_tac simps)]))
+ in
+ fold (fn (s, tvs) => fn thy => AxClass.prove_arity
+ (s, map (inter_sort thy sort o snd) tvs, [cp_class])
+ (Class.intro_classes_tac [] THEN ALLGOALS (resolve_tac thms)) thy)
+ (full_new_type_names' ~~ tyvars) thy
+ end;
+
+ val (perm_thmss,thy3) = thy2 |>
+ fold (fn name1 => fold (composition_instance name1) atoms) atoms |>
+ fold (fn atom => fn thy =>
+ let val pt_name = pt_class_of thy atom
+ in
+ fold (fn (s, tvs) => fn thy => AxClass.prove_arity
+ (s, map (inter_sort thy [pt_name] o snd) tvs, [pt_name])
+ (EVERY
+ [Class.intro_classes_tac [],
+ resolve_tac perm_empty_thms 1,
+ resolve_tac perm_append_thms 1,
+ resolve_tac perm_eq_thms 1, assume_tac 1]) thy)
+ (full_new_type_names' ~~ tyvars) thy
+ end) atoms |>
+ PureThy.add_thmss
+ [((Binding.name (space_implode "_" new_type_names ^ "_unfolded_perm_eq"),
+ unfolded_perm_eq_thms), [Simplifier.simp_add]),
+ ((Binding.name (space_implode "_" new_type_names ^ "_perm_empty"),
+ perm_empty_thms), [Simplifier.simp_add]),
+ ((Binding.name (space_implode "_" new_type_names ^ "_perm_append"),
+ perm_append_thms), [Simplifier.simp_add]),
+ ((Binding.name (space_implode "_" new_type_names ^ "_perm_eq"),
+ perm_eq_thms), [Simplifier.simp_add])];
+
+ (**** Define representing sets ****)
+
+ val _ = warning "representing sets";
+
+ val rep_set_names = DatatypeProp.indexify_names
+ (map (fn (i, _) => name_of_typ (nth_dtyp i) ^ "_set") descr);
+ val big_rep_name =
+ space_implode "_" (DatatypeProp.indexify_names (List.mapPartial
+ (fn (i, ("Nominal.noption", _, _)) => NONE
+ | (i, _) => SOME (name_of_typ (nth_dtyp i))) descr)) ^ "_set";
+ val _ = warning ("big_rep_name: " ^ big_rep_name);
+
+ fun strip_option (dtf as DtType ("fun", [dt, DtRec i])) =
+ (case AList.lookup op = descr i of
+ SOME ("Nominal.noption", _, [(_, [dt']), _]) =>
+ apfst (cons dt) (strip_option dt')
+ | _ => ([], dtf))
+ | strip_option (DtType ("fun", [dt, DtType ("Nominal.noption", [dt'])])) =
+ apfst (cons dt) (strip_option dt')
+ | strip_option dt = ([], dt);
+
+ val dt_atomTs = distinct op = (map (typ_of_dtyp descr sorts)
+ (List.concat (map (fn (_, (_, _, cs)) => List.concat
+ (map (List.concat o map (fst o strip_option) o snd) cs)) descr)));
+ val dt_atoms = map (fst o dest_Type) dt_atomTs;
+
+ fun make_intr s T (cname, cargs) =
+ let
+ fun mk_prem (dt, (j, j', prems, ts)) =
+ let
+ val (dts, dt') = strip_option dt;
+ val (dts', dt'') = strip_dtyp dt';
+ val Ts = map (typ_of_dtyp descr sorts) dts;
+ val Us = map (typ_of_dtyp descr sorts) dts';
+ val T = typ_of_dtyp descr sorts dt'';
+ val free = mk_Free "x" (Us ---> T) j;
+ val free' = app_bnds free (length Us);
+ fun mk_abs_fun (T, (i, t)) =
+ let val U = fastype_of t
+ in (i + 1, Const ("Nominal.abs_fun", [T, U, T] --->
+ Type ("Nominal.noption", [U])) $ mk_Free "y" T i $ t)
+ end
+ in (j + 1, j' + length Ts,
+ case dt'' of
+ DtRec k => list_all (map (pair "x") Us,
+ HOLogic.mk_Trueprop (Free (List.nth (rep_set_names, k),
+ T --> HOLogic.boolT) $ free')) :: prems
+ | _ => prems,
+ snd (List.foldr mk_abs_fun (j', free) Ts) :: ts)
+ end;
+
+ val (_, _, prems, ts) = List.foldr mk_prem (1, 1, [], []) cargs;
+ val concl = HOLogic.mk_Trueprop (Free (s, T --> HOLogic.boolT) $
+ list_comb (Const (cname, map fastype_of ts ---> T), ts))
+ in Logic.list_implies (prems, concl)
+ end;
+
+ val (intr_ts, (rep_set_names', recTs')) =
+ apfst List.concat (apsnd ListPair.unzip (ListPair.unzip (List.mapPartial
+ (fn ((_, ("Nominal.noption", _, _)), _) => NONE
+ | ((i, (_, _, constrs)), rep_set_name) =>
+ let val T = nth_dtyp i
+ in SOME (map (make_intr rep_set_name T) constrs,
+ (rep_set_name, T))
+ end)
+ (descr ~~ rep_set_names))));
+ val rep_set_names'' = map (Sign.full_bname thy3) rep_set_names';
+
+ val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy4) =
+ Inductive.add_inductive_global (serial_string ())
+ {quiet_mode = false, verbose = false, kind = Thm.internalK,
+ alt_name = Binding.name big_rep_name, coind = false, no_elim = true, no_ind = false,
+ skip_mono = true, fork_mono = false}
+ (map (fn (s, T) => ((Binding.name s, T --> HOLogic.boolT), NoSyn))
+ (rep_set_names' ~~ recTs'))
+ [] (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] thy3;
+
+ (**** Prove that representing set is closed under permutation ****)
+
+ val _ = warning "proving closure under permutation...";
+
+ val abs_perm = PureThy.get_thms thy4 "abs_perm";
+
+ val perm_indnames' = List.mapPartial
+ (fn (x, (_, ("Nominal.noption", _, _))) => NONE | (x, _) => SOME x)
+ (perm_indnames ~~ descr);
+
+ fun mk_perm_closed name = map (fn th => standard (th RS mp))
+ (List.take (split_conj_thm (Goal.prove_global thy4 [] []
+ (augment_sort thy4
+ (pt_class_of thy4 name :: map (cp_class_of thy4 name) (dt_atoms \ name))
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map
+ (fn ((s, T), x) =>
+ let
+ val S = Const (s, T --> HOLogic.boolT);
+ val permT = mk_permT (Type (name, []))
+ in HOLogic.mk_imp (S $ Free (x, T),
+ S $ (Const ("Nominal.perm", permT --> T --> T) $
+ Free ("pi", permT) $ Free (x, T)))
+ end) (rep_set_names'' ~~ recTs' ~~ perm_indnames')))))
+ (fn _ => EVERY
+ [indtac rep_induct [] 1,
+ ALLGOALS (simp_tac (simpset_of thy4 addsimps
+ (symmetric perm_fun_def :: abs_perm))),
+ ALLGOALS (resolve_tac rep_intrs THEN_ALL_NEW assume_tac)])),
+ length new_type_names));
+
+ val perm_closed_thmss = map mk_perm_closed atoms;
+
+ (**** typedef ****)
+
+ val _ = warning "defining type...";
+
+ val (typedefs, thy6) =
+ thy4
+ |> fold_map (fn ((((name, mx), tvs), (cname, U)), name') => fn thy =>
+ Typedef.add_typedef false (SOME (Binding.name name'))
+ (Binding.name name, map fst tvs, mx)
+ (Const ("Collect", (U --> HOLogic.boolT) --> HOLogic.mk_setT U) $
+ Const (cname, U --> HOLogic.boolT)) NONE
+ (rtac exI 1 THEN rtac CollectI 1 THEN
+ QUIET_BREADTH_FIRST (has_fewer_prems 1)
+ (resolve_tac rep_intrs 1)) thy |> (fn ((_, r), thy) =>
+ let
+ val permT = mk_permT
+ (TFree (Name.variant (map fst tvs) "'a", HOLogic.typeS));
+ val pi = Free ("pi", permT);
+ val T = Type (Sign.intern_type thy name, map TFree tvs);
+ in apfst (pair r o hd)
+ (PureThy.add_defs_unchecked true [((Binding.name ("prm_" ^ name ^ "_def"), Logic.mk_equals
+ (Const ("Nominal.perm", permT --> T --> T) $ pi $ Free ("x", T),
+ Const (Sign.intern_const thy ("Abs_" ^ name), U --> T) $
+ (Const ("Nominal.perm", permT --> U --> U) $ pi $
+ (Const (Sign.intern_const thy ("Rep_" ^ name), T --> U) $
+ Free ("x", T))))), [])] thy)
+ end))
+ (types_syntax ~~ tyvars ~~
+ List.take (rep_set_names'' ~~ recTs', length new_type_names) ~~
+ new_type_names);
+
+ val perm_defs = map snd typedefs;
+ val Abs_inverse_thms = map (collect_simp o #Abs_inverse o fst) typedefs;
+ val Rep_inverse_thms = map (#Rep_inverse o fst) typedefs;
+ val Rep_thms = map (collect_simp o #Rep o fst) typedefs;
+
+
+ (** prove that new types are in class pt_<name> **)
+
+ val _ = warning "prove that new types are in class pt_<name> ...";
+
+ fun pt_instance (atom, perm_closed_thms) =
+ fold (fn ((((((Abs_inverse, Rep_inverse), Rep),
+ perm_def), name), tvs), perm_closed) => fn thy =>
+ let
+ val pt_class = pt_class_of thy atom;
+ val sort = Sign.certify_sort thy
+ (pt_class :: map (cp_class_of thy atom) (dt_atoms \ atom))
+ in AxClass.prove_arity
+ (Sign.intern_type thy name,
+ map (inter_sort thy sort o snd) tvs, [pt_class])
+ (EVERY [Class.intro_classes_tac [],
+ rewrite_goals_tac [perm_def],
+ asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1,
+ asm_full_simp_tac (simpset_of thy addsimps
+ [Rep RS perm_closed RS Abs_inverse]) 1,
+ asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy
+ ("pt_" ^ Long_Name.base_name atom ^ "3")]) 1]) thy
+ end)
+ (Abs_inverse_thms ~~ Rep_inverse_thms ~~ Rep_thms ~~ perm_defs ~~
+ new_type_names ~~ tyvars ~~ perm_closed_thms);
+
+
+ (** prove that new types are in class cp_<name1>_<name2> **)
+
+ val _ = warning "prove that new types are in class cp_<name1>_<name2> ...";
+
+ fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy =
+ let
+ val cp_class = cp_class_of thy atom1 atom2;
+ val sort = Sign.certify_sort thy
+ (pt_class_of thy atom1 :: map (cp_class_of thy atom1) (dt_atoms \ atom1) @
+ (if atom1 = atom2 then [cp_class_of thy atom1 atom1] else
+ pt_class_of thy atom2 :: map (cp_class_of thy atom2) (dt_atoms \ atom2)));
+ val cp1' = cp_inst_of thy atom1 atom2 RS cp1
+ in fold (fn ((((((Abs_inverse, Rep),
+ perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy =>
+ AxClass.prove_arity
+ (Sign.intern_type thy name,
+ map (inter_sort thy sort o snd) tvs, [cp_class])
+ (EVERY [Class.intro_classes_tac [],
+ rewrite_goals_tac [perm_def],
+ asm_full_simp_tac (simpset_of thy addsimps
+ ((Rep RS perm_closed1 RS Abs_inverse) ::
+ (if atom1 = atom2 then []
+ else [Rep RS perm_closed2 RS Abs_inverse]))) 1,
+ cong_tac 1,
+ rtac refl 1,
+ rtac cp1' 1]) thy)
+ (Abs_inverse_thms ~~ Rep_thms ~~ perm_defs ~~ new_type_names ~~
+ tyvars ~~ perm_closed_thms1 ~~ perm_closed_thms2) thy
+ end;
+
+ val thy7 = fold (fn x => fn thy => thy |>
+ pt_instance x |>
+ fold (cp_instance x) (atoms ~~ perm_closed_thmss))
+ (atoms ~~ perm_closed_thmss) thy6;
+
+ (**** constructors ****)
+
+ fun mk_abs_fun (x, t) =
+ let
+ val T = fastype_of x;
+ val U = fastype_of t
+ in
+ Const ("Nominal.abs_fun", T --> U --> T -->
+ Type ("Nominal.noption", [U])) $ x $ t
+ end;
+
+ val (ty_idxs, _) = List.foldl
+ (fn ((i, ("Nominal.noption", _, _)), p) => p
+ | ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr;
+
+ fun reindex (DtType (s, dts)) = DtType (s, map reindex dts)
+ | reindex (DtRec i) = DtRec (the (AList.lookup op = ty_idxs i))
+ | reindex dt = dt;
+
+ fun strip_suffix i s = implode (List.take (explode s, size s - i));
+
+ (** strips the "_Rep" in type names *)
+ fun strip_nth_name i s =
+ let val xs = Long_Name.explode s;
+ in Long_Name.implode (Library.nth_map (length xs - i) (strip_suffix 4) xs) end;
+
+ val (descr'', ndescr) = ListPair.unzip (map_filter
+ (fn (i, ("Nominal.noption", _, _)) => NONE
+ | (i, (s, dts, constrs)) =>
+ let
+ val SOME index = AList.lookup op = ty_idxs i;
+ val (constrs2, constrs1) =
+ map_split (fn (cname, cargs) =>
+ apsnd (pair (strip_nth_name 2 (strip_nth_name 1 cname)))
+ (fold_map (fn dt => fn dts =>
+ let val (dts', dt') = strip_option dt
+ in ((length dts, length dts'), dts @ dts' @ [reindex dt']) end)
+ cargs [])) constrs
+ in SOME ((index, (strip_nth_name 1 s, map reindex dts, constrs1)),
+ (index, constrs2))
+ end) descr);
+
+ val (descr1, descr2) = chop (length new_type_names) descr'';
+ val descr' = [descr1, descr2];
+
+ fun partition_cargs idxs xs = map (fn (i, j) =>
+ (List.take (List.drop (xs, i), j), List.nth (xs, i + j))) idxs;
+
+ val pdescr = map (fn ((i, (s, dts, constrs)), (_, idxss)) => (i, (s, dts,
+ map (fn ((cname, cargs), idxs) => (cname, partition_cargs idxs cargs))
+ (constrs ~~ idxss)))) (descr'' ~~ ndescr);
+
+ fun nth_dtyp' i = typ_of_dtyp descr'' sorts (DtRec i);
+
+ val rep_names = map (fn s =>
+ Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names;
+ val abs_names = map (fn s =>
+ Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names;
+
+ val recTs = get_rec_types descr'' sorts;
+ val newTs' = Library.take (length new_type_names, recTs');
+ val newTs = Library.take (length new_type_names, recTs);
+
+ val full_new_type_names = map (Sign.full_bname thy) new_type_names;
+
+ fun make_constr_def tname T T' ((thy, defs, eqns),
+ (((cname_rep, _), (cname, cargs)), (cname', mx))) =
+ let
+ fun constr_arg ((dts, dt), (j, l_args, r_args)) =
+ let
+ val xs = map (fn (dt, i) => mk_Free "x" (typ_of_dtyp descr'' sorts dt) i)
+ (dts ~~ (j upto j + length dts - 1))
+ val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
+ in
+ (j + length dts + 1,
+ xs @ x :: l_args,
+ List.foldr mk_abs_fun
+ (case dt of
+ DtRec k => if k < length new_type_names then
+ Const (List.nth (rep_names, k), typ_of_dtyp descr'' sorts dt -->
+ typ_of_dtyp descr sorts dt) $ x
+ else error "nested recursion not (yet) supported"
+ | _ => x) xs :: r_args)
+ end
+
+ val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) cargs;
+ val abs_name = Sign.intern_const thy ("Abs_" ^ tname);
+ val rep_name = Sign.intern_const thy ("Rep_" ^ tname);
+ val constrT = map fastype_of l_args ---> T;
+ val lhs = list_comb (Const (cname, constrT), l_args);
+ val rhs = list_comb (Const (cname_rep, map fastype_of r_args ---> T'), r_args);
+ val def = Logic.mk_equals (lhs, Const (abs_name, T' --> T) $ rhs);
+ val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (Const (rep_name, T --> T') $ lhs, rhs));
+ val def_name = (Long_Name.base_name cname) ^ "_def";
+ val ([def_thm], thy') = thy |>
+ Sign.add_consts_i [(Binding.name cname', constrT, mx)] |>
+ (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)]
+ in (thy', defs @ [def_thm], eqns @ [eqn]) end;
+
+ fun dt_constr_defs ((thy, defs, eqns, dist_lemmas), ((((((_, (_, _, constrs)),
+ (_, (_, _, constrs'))), tname), T), T'), constr_syntax)) =
+ let
+ val rep_const = cterm_of thy
+ (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T'));
+ val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
+ val (thy', defs', eqns') = Library.foldl (make_constr_def tname T T')
+ ((Sign.add_path tname thy, defs, []), constrs ~~ constrs' ~~ constr_syntax)
+ in
+ (parent_path (#flat_names config) thy', defs', eqns @ [eqns'], dist_lemmas @ [dist])
+ end;
+
+ val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = Library.foldl dt_constr_defs
+ ((thy7, [], [], []), List.take (descr, length new_type_names) ~~
+ List.take (pdescr, length new_type_names) ~~
+ new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax);
+
+ val abs_inject_thms = map (collect_simp o #Abs_inject o fst) typedefs
+ val rep_inject_thms = map (#Rep_inject o fst) typedefs
+
+ (* prove theorem Rep_i (Constr_j ...) = Constr'_j ... *)
+
+ fun prove_constr_rep_thm eqn =
+ let
+ val inj_thms = map (fn r => r RS iffD1) abs_inject_thms;
+ val rewrites = constr_defs @ map mk_meta_eq Rep_inverse_thms
+ in Goal.prove_global thy8 [] [] eqn (fn _ => EVERY
+ [resolve_tac inj_thms 1,
+ rewrite_goals_tac rewrites,
+ rtac refl 3,
+ resolve_tac rep_intrs 2,
+ REPEAT (resolve_tac Rep_thms 1)])
+ end;
+
+ val constr_rep_thmss = map (map prove_constr_rep_thm) constr_rep_eqns;
+
+ (* prove theorem pi \<bullet> Rep_i x = Rep_i (pi \<bullet> x) *)
+
+ fun prove_perm_rep_perm (atom, perm_closed_thms) = map (fn th =>
+ let
+ val _ $ (_ $ (Rep $ x)) = Logic.unvarify (prop_of th);
+ val Type ("fun", [T, U]) = fastype_of Rep;
+ val permT = mk_permT (Type (atom, []));
+ val pi = Free ("pi", permT);
+ in
+ Goal.prove_global thy8 [] []
+ (augment_sort thy8
+ (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (dt_atoms \ atom))
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (Const ("Nominal.perm", permT --> U --> U) $ pi $ (Rep $ x),
+ Rep $ (Const ("Nominal.perm", permT --> T --> T) $ pi $ x)))))
+ (fn _ => simp_tac (HOL_basic_ss addsimps (perm_defs @ Abs_inverse_thms @
+ perm_closed_thms @ Rep_thms)) 1)
+ end) Rep_thms;
+
+ val perm_rep_perm_thms = List.concat (map prove_perm_rep_perm
+ (atoms ~~ perm_closed_thmss));
+
+ (* prove distinctness theorems *)
+
+ val distinct_props = DatatypeProp.make_distincts descr' sorts;
+ val dist_rewrites = map2 (fn rep_thms => fn dist_lemma =>
+ dist_lemma :: rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0])
+ constr_rep_thmss dist_lemmas;
+
+ fun prove_distinct_thms _ (_, []) = []
+ | prove_distinct_thms (p as (rep_thms, dist_lemma)) (k, t :: ts) =
+ let
+ val dist_thm = Goal.prove_global thy8 [] [] t (fn _ =>
+ simp_tac (simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1)
+ in dist_thm :: standard (dist_thm RS not_sym) ::
+ prove_distinct_thms p (k, ts)
+ end;
+
+ val distinct_thms = map2 prove_distinct_thms
+ (constr_rep_thmss ~~ dist_lemmas) distinct_props;
+
+ (** prove equations for permutation functions **)
+
+ val perm_simps' = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
+ let val T = nth_dtyp' i
+ in List.concat (map (fn (atom, perm_closed_thms) =>
+ map (fn ((cname, dts), constr_rep_thm) =>
+ let
+ val cname = Sign.intern_const thy8
+ (Long_Name.append tname (Long_Name.base_name cname));
+ val permT = mk_permT (Type (atom, []));
+ val pi = Free ("pi", permT);
+
+ fun perm t =
+ let val T = fastype_of t
+ in Const ("Nominal.perm", permT --> T --> T) $ pi $ t end;
+
+ fun constr_arg ((dts, dt), (j, l_args, r_args)) =
+ let
+ val Ts = map (typ_of_dtyp descr'' sorts) dts;
+ val xs = map (fn (T, i) => mk_Free "x" T i)
+ (Ts ~~ (j upto j + length dts - 1))
+ val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
+ in
+ (j + length dts + 1,
+ xs @ x :: l_args,
+ map perm (xs @ [x]) @ r_args)
+ end
+
+ val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) dts;
+ val c = Const (cname, map fastype_of l_args ---> T)
+ in
+ Goal.prove_global thy8 [] []
+ (augment_sort thy8
+ (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (dt_atoms \ atom))
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (perm (list_comb (c, l_args)), list_comb (c, r_args)))))
+ (fn _ => EVERY
+ [simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1,
+ simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @
+ constr_defs @ perm_closed_thms)) 1,
+ TRY (simp_tac (HOL_basic_ss addsimps
+ (symmetric perm_fun_def :: abs_perm)) 1),
+ TRY (simp_tac (HOL_basic_ss addsimps
+ (perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @
+ perm_closed_thms)) 1)])
+ end) (constrs ~~ constr_rep_thms)) (atoms ~~ perm_closed_thmss))
+ end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
+
+ (** prove injectivity of constructors **)
+
+ val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms;
+ val alpha = PureThy.get_thms thy8 "alpha";
+ val abs_fresh = PureThy.get_thms thy8 "abs_fresh";
+
+ val pt_cp_sort =
+ map (pt_class_of thy8) dt_atoms @
+ maps (fn s => map (cp_class_of thy8 s) (dt_atoms \ s)) dt_atoms;
+
+ val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
+ let val T = nth_dtyp' i
+ in List.mapPartial (fn ((cname, dts), constr_rep_thm) =>
+ if null dts then NONE else SOME
+ let
+ val cname = Sign.intern_const thy8
+ (Long_Name.append tname (Long_Name.base_name cname));
+
+ fun make_inj ((dts, dt), (j, args1, args2, eqs)) =
+ let
+ val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1);
+ val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
+ val ys = map (fn (T, i) => mk_Free "y" T i) Ts_idx;
+ val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts);
+ val y = mk_Free "y" (typ_of_dtyp descr'' sorts dt) (j + length dts)
+ in
+ (j + length dts + 1,
+ xs @ (x :: args1), ys @ (y :: args2),
+ HOLogic.mk_eq
+ (List.foldr mk_abs_fun x xs, List.foldr mk_abs_fun y ys) :: eqs)
+ end;
+
+ val (_, args1, args2, eqs) = List.foldr make_inj (1, [], [], []) dts;
+ val Ts = map fastype_of args1;
+ val c = Const (cname, Ts ---> T)
+ in
+ Goal.prove_global thy8 [] []
+ (augment_sort thy8 pt_cp_sort
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)),
+ foldr1 HOLogic.mk_conj eqs))))
+ (fn _ => EVERY
+ [asm_full_simp_tac (simpset_of thy8 addsimps (constr_rep_thm ::
+ rep_inject_thms')) 1,
+ TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def ::
+ alpha @ abs_perm @ abs_fresh @ rep_inject_thms @
+ perm_rep_perm_thms)) 1)])
+ end) (constrs ~~ constr_rep_thms)
+ end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
+
+ (** equations for support and freshness **)
+
+ val (supp_thms, fresh_thms) = ListPair.unzip (map ListPair.unzip
+ (map (fn ((((i, (_, _, constrs)), tname), inject_thms'), perm_thms') =>
+ let val T = nth_dtyp' i
+ in List.concat (map (fn (cname, dts) => map (fn atom =>
+ let
+ val cname = Sign.intern_const thy8
+ (Long_Name.append tname (Long_Name.base_name cname));
+ val atomT = Type (atom, []);
+
+ fun process_constr ((dts, dt), (j, args1, args2)) =
+ let
+ val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1);
+ val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
+ val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
+ in
+ (j + length dts + 1,
+ xs @ (x :: args1), List.foldr mk_abs_fun x xs :: args2)
+ end;
+
+ val (_, args1, args2) = List.foldr process_constr (1, [], []) dts;
+ val Ts = map fastype_of args1;
+ val c = list_comb (Const (cname, Ts ---> T), args1);
+ fun supp t =
+ Const ("Nominal.supp", fastype_of t --> HOLogic.mk_setT atomT) $ t;
+ fun fresh t = fresh_const atomT (fastype_of t) $ Free ("a", atomT) $ t;
+ val supp_thm = Goal.prove_global thy8 [] []
+ (augment_sort thy8 pt_cp_sort
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (supp c,
+ if null dts then HOLogic.mk_set atomT []
+ else foldr1 (HOLogic.mk_binop @{const_name Un}) (map supp args2)))))
+ (fn _ =>
+ simp_tac (HOL_basic_ss addsimps (supp_def ::
+ Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un ::
+ symmetric empty_def :: finite_emptyI :: simp_thms @
+ abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1)
+ in
+ (supp_thm,
+ Goal.prove_global thy8 [] [] (augment_sort thy8 pt_cp_sort
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (fresh c,
+ if null dts then HOLogic.true_const
+ else foldr1 HOLogic.mk_conj (map fresh args2)))))
+ (fn _ =>
+ simp_tac (HOL_ss addsimps [Un_iff, empty_iff, fresh_def, supp_thm]) 1))
+ end) atoms) constrs)
+ end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps')));
+
+ (**** weak induction theorem ****)
+
+ fun mk_indrule_lemma ((prems, concls), (((i, _), T), U)) =
+ let
+ val Rep_t = Const (List.nth (rep_names, i), T --> U) $
+ mk_Free "x" T i;
+
+ val Abs_t = Const (List.nth (abs_names, i), U --> T)
+
+ in (prems @ [HOLogic.imp $
+ (Const (List.nth (rep_set_names'', i), U --> HOLogic.boolT) $ Rep_t) $
+ (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
+ concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i])
+ end;
+
+ val (indrule_lemma_prems, indrule_lemma_concls) =
+ Library.foldl mk_indrule_lemma (([], []), (descr'' ~~ recTs ~~ recTs'));
+
+ val indrule_lemma = Goal.prove_global thy8 [] []
+ (Logic.mk_implies
+ (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems),
+ HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY
+ [REPEAT (etac conjE 1),
+ REPEAT (EVERY
+ [TRY (rtac conjI 1), full_simp_tac (HOL_basic_ss addsimps Rep_inverse_thms) 1,
+ etac mp 1, resolve_tac Rep_thms 1])]);
+
+ val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
+ val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
+ map (Free o apfst fst o dest_Var) Ps;
+ val indrule_lemma' = cterm_instantiate
+ (map (cterm_of thy8) Ps ~~ map (cterm_of thy8) frees) indrule_lemma;
+
+ val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms;
+
+ val dt_induct_prop = DatatypeProp.make_ind descr' sorts;
+ val dt_induct = Goal.prove_global thy8 []
+ (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
+ (fn {prems, ...} => EVERY
+ [rtac indrule_lemma' 1,
+ (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
+ EVERY (map (fn (prem, r) => (EVERY
+ [REPEAT (eresolve_tac Abs_inverse_thms' 1),
+ simp_tac (HOL_basic_ss addsimps [symmetric r]) 1,
+ DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
+ (prems ~~ constr_defs))]);
+
+ val case_names_induct = mk_case_names_induct descr'';
+
+ (**** prove that new datatypes have finite support ****)
+
+ val _ = warning "proving finite support for the new datatype";
+
+ val indnames = DatatypeProp.make_tnames recTs;
+
+ val abs_supp = PureThy.get_thms thy8 "abs_supp";
+ val supp_atm = PureThy.get_thms thy8 "supp_atm";
+
+ val finite_supp_thms = map (fn atom =>
+ let val atomT = Type (atom, [])
+ in map standard (List.take
+ (split_conj_thm (Goal.prove_global thy8 [] []
+ (augment_sort thy8 (fs_class_of thy8 atom :: pt_cp_sort)
+ (HOLogic.mk_Trueprop
+ (foldr1 HOLogic.mk_conj (map (fn (s, T) =>
+ Const ("Finite_Set.finite", HOLogic.mk_setT atomT --> HOLogic.boolT) $
+ (Const ("Nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T)))
+ (indnames ~~ recTs)))))
+ (fn _ => indtac dt_induct indnames 1 THEN
+ ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps
+ (abs_supp @ supp_atm @
+ PureThy.get_thms thy8 ("fs_" ^ Long_Name.base_name atom ^ "1") @
+ List.concat supp_thms))))),
+ length new_type_names))
+ end) atoms;
+
+ val simp_atts = replicate (length new_type_names) [Simplifier.simp_add];
+
+ (* Function to add both the simp and eqvt attributes *)
+ (* These two attributes are duplicated on all the types in the mutual nominal datatypes *)
+
+ val simp_eqvt_atts = replicate (length new_type_names) [Simplifier.simp_add, NominalThmDecls.eqvt_add];
+
+ val (_, thy9) = thy8 |>
+ Sign.add_path big_name |>
+ PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] ||>>
+ PureThy.add_thmss [((Binding.name "inducts", projections dt_induct), [case_names_induct])] ||>
+ Sign.parent_path ||>>
+ DatatypeAux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>>
+ DatatypeAux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>>
+ DatatypeAux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>>
+ DatatypeAux.store_thmss "inject" new_type_names inject_thms ||>>
+ DatatypeAux.store_thmss "supp" new_type_names supp_thms ||>>
+ DatatypeAux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||>
+ fold (fn (atom, ths) => fn thy =>
+ let
+ val class = fs_class_of thy atom;
+ val sort = Sign.certify_sort thy (class :: pt_cp_sort)
+ in fold (fn Type (s, Ts) => AxClass.prove_arity
+ (s, map (inter_sort thy sort o snd o dest_TFree) Ts, [class])
+ (Class.intro_classes_tac [] THEN resolve_tac ths 1)) newTs thy
+ end) (atoms ~~ finite_supp_thms);
+
+ (**** strong induction theorem ****)
+
+ val pnames = if length descr'' = 1 then ["P"]
+ else map (fn i => "P" ^ string_of_int i) (1 upto length descr'');
+ val ind_sort = if null dt_atomTs then HOLogic.typeS
+ else Sign.certify_sort thy9 (map (fs_class_of thy9) dt_atoms);
+ val fsT = TFree ("'n", ind_sort);
+ val fsT' = TFree ("'n", HOLogic.typeS);
+
+ val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T)))
+ (DatatypeProp.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs);
+
+ fun make_pred fsT i T =
+ Free (List.nth (pnames, i), fsT --> T --> HOLogic.boolT);
+
+ fun mk_fresh1 xs [] = []
+ | mk_fresh1 xs ((y as (_, T)) :: ys) = map (fn x => HOLogic.mk_Trueprop
+ (HOLogic.mk_not (HOLogic.mk_eq (Free y, Free x))))
+ (filter (fn (_, U) => T = U) (rev xs)) @
+ mk_fresh1 (y :: xs) ys;
+
+ fun mk_fresh2 xss [] = []
+ | mk_fresh2 xss ((p as (ys, _)) :: yss) = List.concat (map (fn y as (_, T) =>
+ map (fn (_, x as (_, U)) => HOLogic.mk_Trueprop
+ (fresh_const T U $ Free y $ Free x)) (rev xss @ yss)) ys) @
+ mk_fresh2 (p :: xss) yss;
+
+ fun make_ind_prem fsT f k T ((cname, cargs), idxs) =
+ let
+ val recs = List.filter is_rec_type cargs;
+ val Ts = map (typ_of_dtyp descr'' sorts) cargs;
+ val recTs' = map (typ_of_dtyp descr'' sorts) recs;
+ val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts);
+ val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
+ val frees = tnames ~~ Ts;
+ val frees' = partition_cargs idxs frees;
+ val z = (Name.variant tnames "z", fsT);
+
+ fun mk_prem ((dt, s), T) =
+ let
+ val (Us, U) = strip_type T;
+ val l = length Us
+ in list_all (z :: map (pair "x") Us, HOLogic.mk_Trueprop
+ (make_pred fsT (body_index dt) U $ Bound l $ app_bnds (Free (s, T)) l))
+ end;
+
+ val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs');
+ val prems' = map (fn p as (_, T) => HOLogic.mk_Trueprop
+ (f T (Free p) (Free z))) (List.concat (map fst frees')) @
+ mk_fresh1 [] (List.concat (map fst frees')) @
+ mk_fresh2 [] frees'
+
+ in list_all_free (frees @ [z], Logic.list_implies (prems' @ prems,
+ HOLogic.mk_Trueprop (make_pred fsT k T $ Free z $
+ list_comb (Const (cname, Ts ---> T), map Free frees))))
+ end;
+
+ val ind_prems = List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) =>
+ map (make_ind_prem fsT (fn T => fn t => fn u =>
+ fresh_const T fsT $ t $ u) i T)
+ (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs));
+ val tnames = DatatypeProp.make_tnames recTs;
+ val zs = Name.variant_list tnames (replicate (length descr'') "z");
+ val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+ (map (fn ((((i, _), T), tname), z) =>
+ make_pred fsT i T $ Free (z, fsT) $ Free (tname, T))
+ (descr'' ~~ recTs ~~ tnames ~~ zs)));
+ val induct = Logic.list_implies (ind_prems, ind_concl);
+
+ val ind_prems' =
+ map (fn (_, f as Free (_, T)) => list_all_free ([("x", fsT')],
+ HOLogic.mk_Trueprop (Const ("Finite_Set.finite",
+ (snd (split_last (binder_types T)) --> HOLogic.boolT) -->
+ HOLogic.boolT) $ (f $ Free ("x", fsT'))))) fresh_fs @
+ List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) =>
+ map (make_ind_prem fsT' (fn T => fn t => fn u => HOLogic.Not $
+ HOLogic.mk_mem (t, the (AList.lookup op = fresh_fs T) $ u)) i T)
+ (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs));
+ val ind_concl' = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+ (map (fn ((((i, _), T), tname), z) =>
+ make_pred fsT' i T $ Free (z, fsT') $ Free (tname, T))
+ (descr'' ~~ recTs ~~ tnames ~~ zs)));
+ val induct' = Logic.list_implies (ind_prems', ind_concl');
+
+ val aux_ind_vars =
+ (DatatypeProp.indexify_names (replicate (length dt_atomTs) "pi") ~~
+ map mk_permT dt_atomTs) @ [("z", fsT')];
+ val aux_ind_Ts = rev (map snd aux_ind_vars);
+ val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+ (map (fn (((i, _), T), tname) =>
+ HOLogic.list_all (aux_ind_vars, make_pred fsT' i T $ Bound 0 $
+ fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1))
+ (Free (tname, T))))
+ (descr'' ~~ recTs ~~ tnames)));
+
+ val fin_set_supp = map (fn s =>
+ at_inst_of thy9 s RS at_fin_set_supp) dt_atoms;
+ val fin_set_fresh = map (fn s =>
+ at_inst_of thy9 s RS at_fin_set_fresh) dt_atoms;
+ val pt1_atoms = map (fn Type (s, _) =>
+ PureThy.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "1")) dt_atomTs;
+ val pt2_atoms = map (fn Type (s, _) =>
+ PureThy.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "2") RS sym) dt_atomTs;
+ val exists_fresh' = PureThy.get_thms thy9 "exists_fresh'";
+ val fs_atoms = PureThy.get_thms thy9 "fin_supp";
+ val abs_supp = PureThy.get_thms thy9 "abs_supp";
+ val perm_fresh_fresh = PureThy.get_thms thy9 "perm_fresh_fresh";
+ val calc_atm = PureThy.get_thms thy9 "calc_atm";
+ val fresh_atm = PureThy.get_thms thy9 "fresh_atm";
+ val fresh_left = PureThy.get_thms thy9 "fresh_left";
+ val perm_swap = PureThy.get_thms thy9 "perm_swap";
+
+ fun obtain_fresh_name' ths ts T (freshs1, freshs2, ctxt) =
+ let
+ val p = foldr1 HOLogic.mk_prod (ts @ freshs1);
+ val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
+ (HOLogic.exists_const T $ Abs ("x", T,
+ fresh_const T (fastype_of p) $
+ Bound 0 $ p)))
+ (fn _ => EVERY
+ [resolve_tac exists_fresh' 1,
+ simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms @
+ fin_set_supp @ ths)) 1]);
+ val (([cx], ths), ctxt') = Obtain.result
+ (fn _ => EVERY
+ [etac exE 1,
+ full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1,
+ REPEAT (etac conjE 1)])
+ [ex] ctxt
+ in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end;
+
+ fun fresh_fresh_inst thy a b =
+ let
+ val T = fastype_of a;
+ val SOME th = find_first (fn th => case prop_of th of
+ _ $ (_ $ (Const (_, Type (_, [U, _])) $ _ $ _)) $ _ => U = T
+ | _ => false) perm_fresh_fresh
+ in
+ Drule.instantiate' []
+ [SOME (cterm_of thy a), NONE, SOME (cterm_of thy b)] th
+ end;
+
+ val fs_cp_sort =
+ map (fs_class_of thy9) dt_atoms @
+ maps (fn s => map (cp_class_of thy9 s) (dt_atoms \ s)) dt_atoms;
+
+ (**********************************************************************
+ The subgoals occurring in the proof of induct_aux have the
+ following parameters:
+
+ x_1 ... x_k p_1 ... p_m z
+
+ where
+
+ x_i : constructor arguments (introduced by weak induction rule)
+ p_i : permutations (one for each atom type in the data type)
+ z : freshness context
+ ***********************************************************************)
+
+ val _ = warning "proving strong induction theorem ...";
+
+ val induct_aux = Goal.prove_global thy9 []
+ (map (augment_sort thy9 fs_cp_sort) ind_prems')
+ (augment_sort thy9 fs_cp_sort ind_concl') (fn {prems, context} =>
+ let
+ val (prems1, prems2) = chop (length dt_atomTs) prems;
+ val ind_ss2 = HOL_ss addsimps
+ finite_Diff :: abs_fresh @ abs_supp @ fs_atoms;
+ val ind_ss1 = ind_ss2 addsimps fresh_left @ calc_atm @
+ fresh_atm @ rev_simps @ app_simps;
+ val ind_ss3 = HOL_ss addsimps abs_fun_eq1 ::
+ abs_perm @ calc_atm @ perm_swap;
+ val ind_ss4 = HOL_basic_ss addsimps fresh_left @ prems1 @
+ fin_set_fresh @ calc_atm;
+ val ind_ss5 = HOL_basic_ss addsimps pt1_atoms;
+ val ind_ss6 = HOL_basic_ss addsimps flat perm_simps';
+ val th = Goal.prove context [] []
+ (augment_sort thy9 fs_cp_sort aux_ind_concl)
+ (fn {context = context1, ...} =>
+ EVERY (indtac dt_induct tnames 1 ::
+ maps (fn ((_, (_, _, constrs)), (_, constrs')) =>
+ map (fn ((cname, cargs), is) =>
+ REPEAT (rtac allI 1) THEN
+ SUBPROOF (fn {prems = iprems, params, concl,
+ context = context2, ...} =>
+ let
+ val concl' = term_of concl;
+ val _ $ (_ $ _ $ u) = concl';
+ val U = fastype_of u;
+ val (xs, params') =
+ chop (length cargs) (map term_of params);
+ val Ts = map fastype_of xs;
+ val cnstr = Const (cname, Ts ---> U);
+ val (pis, z) = split_last params';
+ val mk_pi = fold_rev (mk_perm []) pis;
+ val xs' = partition_cargs is xs;
+ val xs'' = map (fn (ts, u) => (map mk_pi ts, mk_pi u)) xs';
+ val ts = maps (fn (ts, u) => ts @ [u]) xs'';
+ val (freshs1, freshs2, context3) = fold (fn t =>
+ let val T = fastype_of t
+ in obtain_fresh_name' prems1
+ (the (AList.lookup op = fresh_fs T) $ z :: ts) T
+ end) (maps fst xs') ([], [], context2);
+ val freshs1' = unflat (map fst xs') freshs1;
+ val freshs2' = map (Simplifier.simplify ind_ss4)
+ (mk_not_sym freshs2);
+ val ind_ss1' = ind_ss1 addsimps freshs2';
+ val ind_ss3' = ind_ss3 addsimps freshs2';
+ val rename_eq =
+ if forall (null o fst) xs' then []
+ else [Goal.prove context3 [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (list_comb (cnstr, ts),
+ list_comb (cnstr, maps (fn ((bs, t), cs) =>
+ cs @ [fold_rev (mk_perm []) (map perm_of_pair
+ (bs ~~ cs)) t]) (xs'' ~~ freshs1')))))
+ (fn _ => EVERY
+ (simp_tac (HOL_ss addsimps flat inject_thms) 1 ::
+ REPEAT (FIRSTGOAL (rtac conjI)) ::
+ maps (fn ((bs, t), cs) =>
+ if null bs then []
+ else rtac sym 1 :: maps (fn (b, c) =>
+ [rtac trans 1, rtac sym 1,
+ rtac (fresh_fresh_inst thy9 b c) 1,
+ simp_tac ind_ss1' 1,
+ simp_tac ind_ss2 1,
+ simp_tac ind_ss3' 1]) (bs ~~ cs))
+ (xs'' ~~ freshs1')))];
+ val th = Goal.prove context3 [] [] concl' (fn _ => EVERY
+ [simp_tac (ind_ss6 addsimps rename_eq) 1,
+ cut_facts_tac iprems 1,
+ (resolve_tac prems THEN_ALL_NEW
+ SUBGOAL (fn (t, i) => case Logic.strip_assums_concl t of
+ _ $ (Const ("Nominal.fresh", _) $ _ $ _) =>
+ simp_tac ind_ss1' i
+ | _ $ (Const ("Not", _) $ _) =>
+ resolve_tac freshs2' i
+ | _ => asm_simp_tac (HOL_basic_ss addsimps
+ pt2_atoms addsimprocs [perm_simproc]) i)) 1])
+ val final = ProofContext.export context3 context2 [th]
+ in
+ resolve_tac final 1
+ end) context1 1) (constrs ~~ constrs')) (descr'' ~~ ndescr)))
+ in
+ EVERY
+ [cut_facts_tac [th] 1,
+ REPEAT (eresolve_tac [conjE, @{thm allE_Nil}] 1),
+ REPEAT (etac allE 1),
+ REPEAT (TRY (rtac conjI 1) THEN asm_full_simp_tac ind_ss5 1)]
+ end);
+
+ val induct_aux' = Thm.instantiate ([],
+ map (fn (s, v as Var (_, T)) =>
+ (cterm_of thy9 v, cterm_of thy9 (Free (s, T))))
+ (pnames ~~ map head_of (HOLogic.dest_conj
+ (HOLogic.dest_Trueprop (concl_of induct_aux)))) @
+ map (fn (_, f) =>
+ let val f' = Logic.varify f
+ in (cterm_of thy9 f',
+ cterm_of thy9 (Const ("Nominal.supp", fastype_of f')))
+ end) fresh_fs) induct_aux;
+
+ val induct = Goal.prove_global thy9 []
+ (map (augment_sort thy9 fs_cp_sort) ind_prems)
+ (augment_sort thy9 fs_cp_sort ind_concl)
+ (fn {prems, ...} => EVERY
+ [rtac induct_aux' 1,
+ REPEAT (resolve_tac fs_atoms 1),
+ REPEAT ((resolve_tac prems THEN_ALL_NEW
+ (etac meta_spec ORELSE' full_simp_tac (HOL_basic_ss addsimps [fresh_def]))) 1)])
+
+ val (_, thy10) = thy9 |>
+ Sign.add_path big_name |>
+ PureThy.add_thms [((Binding.name "strong_induct'", induct_aux), [])] ||>>
+ PureThy.add_thms [((Binding.name "strong_induct", induct), [case_names_induct])] ||>>
+ PureThy.add_thmss [((Binding.name "strong_inducts", projections induct), [case_names_induct])];
+
+ (**** recursion combinator ****)
+
+ val _ = warning "defining recursion combinator ...";
+
+ val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
+
+ val (rec_result_Ts', rec_fn_Ts') = DatatypeProp.make_primrec_Ts descr' sorts used;
+
+ val rec_sort = if null dt_atomTs then HOLogic.typeS else
+ Sign.certify_sort thy10 pt_cp_sort;
+
+ val rec_result_Ts = map (fn TFree (s, _) => TFree (s, rec_sort)) rec_result_Ts';
+ val rec_fn_Ts = map (typ_subst_atomic (rec_result_Ts' ~~ rec_result_Ts)) rec_fn_Ts';
+
+ val rec_set_Ts = map (fn (T1, T2) =>
+ rec_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
+
+ val big_rec_name = big_name ^ "_rec_set";
+ val rec_set_names' =
+ if length descr'' = 1 then [big_rec_name] else
+ map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int)
+ (1 upto (length descr''));
+ val rec_set_names = map (Sign.full_bname thy10) rec_set_names';
+
+ val rec_fns = map (uncurry (mk_Free "f"))
+ (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts)));
+ val rec_sets' = map (fn c => list_comb (Free c, rec_fns))
+ (rec_set_names' ~~ rec_set_Ts);
+ val rec_sets = map (fn c => list_comb (Const c, rec_fns))
+ (rec_set_names ~~ rec_set_Ts);
+
+ (* introduction rules for graph of recursion function *)
+
+ val rec_preds = map (fn (a, T) =>
+ Free (a, T --> HOLogic.boolT)) (pnames ~~ rec_result_Ts);
+
+ fun mk_fresh3 rs [] = []
+ | mk_fresh3 rs ((p as (ys, z)) :: yss) = List.concat (map (fn y as (_, T) =>
+ List.mapPartial (fn ((_, (_, x)), r as (_, U)) => if z = x then NONE
+ else SOME (HOLogic.mk_Trueprop
+ (fresh_const T U $ Free y $ Free r))) rs) ys) @
+ mk_fresh3 rs yss;
+
+ (* FIXME: avoid collisions with other variable names? *)
+ val rec_ctxt = Free ("z", fsT');
+
+ fun make_rec_intr T p rec_set ((rec_intr_ts, rec_prems, rec_prems',
+ rec_eq_prems, l), ((cname, cargs), idxs)) =
+ let
+ val Ts = map (typ_of_dtyp descr'' sorts) cargs;
+ val frees = map (fn i => "x" ^ string_of_int i) (1 upto length Ts) ~~ Ts;
+ val frees' = partition_cargs idxs frees;
+ val binders = List.concat (map fst frees');
+ val atomTs = distinct op = (maps (map snd o fst) frees');
+ val recs = List.mapPartial
+ (fn ((_, DtRec i), p) => SOME (i, p) | _ => NONE)
+ (partition_cargs idxs cargs ~~ frees');
+ val frees'' = map (fn i => "y" ^ string_of_int i) (1 upto length recs) ~~
+ map (fn (i, _) => List.nth (rec_result_Ts, i)) recs;
+ val prems1 = map (fn ((i, (_, x)), y) => HOLogic.mk_Trueprop
+ (List.nth (rec_sets', i) $ Free x $ Free y)) (recs ~~ frees'');
+ val prems2 =
+ map (fn f => map (fn p as (_, T) => HOLogic.mk_Trueprop
+ (fresh_const T (fastype_of f) $ Free p $ f)) binders) rec_fns;
+ val prems3 = mk_fresh1 [] binders @ mk_fresh2 [] frees';
+ val prems4 = map (fn ((i, _), y) =>
+ HOLogic.mk_Trueprop (List.nth (rec_preds, i) $ Free y)) (recs ~~ frees'');
+ val prems5 = mk_fresh3 (recs ~~ frees'') frees';
+ val prems6 = maps (fn aT => map (fn y as (_, T) => HOLogic.mk_Trueprop
+ (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
+ (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ Free y)))
+ frees'') atomTs;
+ val prems7 = map (fn x as (_, T) => HOLogic.mk_Trueprop
+ (fresh_const T fsT' $ Free x $ rec_ctxt)) binders;
+ val result = list_comb (List.nth (rec_fns, l), map Free (frees @ frees''));
+ val result_freshs = map (fn p as (_, T) =>
+ fresh_const T (fastype_of result) $ Free p $ result) binders;
+ val P = HOLogic.mk_Trueprop (p $ result)
+ in
+ (rec_intr_ts @ [Logic.list_implies (List.concat prems2 @ prems3 @ prems1,
+ HOLogic.mk_Trueprop (rec_set $
+ list_comb (Const (cname, Ts ---> T), map Free frees) $ result))],
+ rec_prems @ [list_all_free (frees @ frees'', Logic.list_implies (prems4, P))],
+ rec_prems' @ map (fn fr => list_all_free (frees @ frees'',
+ Logic.list_implies (List.nth (prems2, l) @ prems3 @ prems5 @ prems7 @ prems6 @ [P],
+ HOLogic.mk_Trueprop fr))) result_freshs,
+ rec_eq_prems @ [List.concat prems2 @ prems3],
+ l + 1)
+ end;
+
+ val (rec_intr_ts, rec_prems, rec_prems', rec_eq_prems, _) =
+ Library.foldl (fn (x, ((((d, d'), T), p), rec_set)) =>
+ Library.foldl (make_rec_intr T p rec_set) (x, #3 (snd d) ~~ snd d'))
+ (([], [], [], [], 0), descr'' ~~ ndescr ~~ recTs ~~ rec_preds ~~ rec_sets');
+
+ val ({intrs = rec_intrs, elims = rec_elims, raw_induct = rec_induct, ...}, thy11) =
+ thy10 |>
+ Inductive.add_inductive_global (serial_string ())
+ {quiet_mode = #quiet config, verbose = false, kind = Thm.internalK,
+ alt_name = Binding.name big_rec_name, coind = false, no_elim = false, no_ind = false,
+ skip_mono = true, fork_mono = false}
+ (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
+ (map dest_Free rec_fns)
+ (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] ||>
+ PureThy.hide_fact true (Long_Name.append (Sign.full_bname thy10 big_rec_name) "induct");
+
+ (** equivariance **)
+
+ val fresh_bij = PureThy.get_thms thy11 "fresh_bij";
+ val perm_bij = PureThy.get_thms thy11 "perm_bij";
+
+ val (rec_equiv_thms, rec_equiv_thms') = ListPair.unzip (map (fn aT =>
+ let
+ val permT = mk_permT aT;
+ val pi = Free ("pi", permT);
+ val rec_fns_pi = map (mk_perm [] pi o uncurry (mk_Free "f"))
+ (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts)));
+ val rec_sets_pi = map (fn c => list_comb (Const c, rec_fns_pi))
+ (rec_set_names ~~ rec_set_Ts);
+ val ps = map (fn ((((T, U), R), R'), i) =>
+ let
+ val x = Free ("x" ^ string_of_int i, T);
+ val y = Free ("y" ^ string_of_int i, U)
+ in
+ (R $ x $ y, R' $ mk_perm [] pi x $ mk_perm [] pi y)
+ end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ rec_sets_pi ~~ (1 upto length recTs));
+ val ths = map (fn th => standard (th RS mp)) (split_conj_thm
+ (Goal.prove_global thy11 [] []
+ (augment_sort thy1 pt_cp_sort
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps))))
+ (fn _ => rtac rec_induct 1 THEN REPEAT
+ (simp_tac (Simplifier.theory_context thy11 HOL_basic_ss
+ addsimps flat perm_simps'
+ addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
+ (resolve_tac rec_intrs THEN_ALL_NEW
+ asm_simp_tac (HOL_ss addsimps (fresh_bij @ perm_bij))) 1))))
+ val ths' = map (fn ((P, Q), th) =>
+ Goal.prove_global thy11 [] []
+ (augment_sort thy1 pt_cp_sort
+ (Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P)))
+ (fn _ => dtac (Thm.instantiate ([],
+ [(cterm_of thy11 (Var (("pi", 0), permT)),
+ cterm_of thy11 (Const ("List.rev", permT --> permT) $ pi))]) th) 1 THEN
+ NominalPermeq.perm_simp_tac HOL_ss 1)) (ps ~~ ths)
+ in (ths, ths') end) dt_atomTs);
+
+ (** finite support **)
+
+ val rec_fin_supp_thms = map (fn aT =>
+ let
+ val name = Long_Name.base_name (fst (dest_Type aT));
+ val fs_name = PureThy.get_thm thy11 ("fs_" ^ name ^ "1");
+ val aset = HOLogic.mk_setT aT;
+ val finite = Const ("Finite_Set.finite", aset --> HOLogic.boolT);
+ val fins = map (fn (f, T) => HOLogic.mk_Trueprop
+ (finite $ (Const ("Nominal.supp", T --> aset) $ f)))
+ (rec_fns ~~ rec_fn_Ts)
+ in
+ map (fn th => standard (th RS mp)) (split_conj_thm
+ (Goal.prove_global thy11 []
+ (map (augment_sort thy11 fs_cp_sort) fins)
+ (augment_sort thy11 fs_cp_sort
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+ (map (fn (((T, U), R), i) =>
+ let
+ val x = Free ("x" ^ string_of_int i, T);
+ val y = Free ("y" ^ string_of_int i, U)
+ in
+ HOLogic.mk_imp (R $ x $ y,
+ finite $ (Const ("Nominal.supp", U --> aset) $ y))
+ end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~
+ (1 upto length recTs))))))
+ (fn {prems = fins, ...} =>
+ (rtac rec_induct THEN_ALL_NEW cut_facts_tac fins) 1 THEN REPEAT
+ (NominalPermeq.finite_guess_tac (HOL_ss addsimps [fs_name]) 1))))
+ end) dt_atomTs;
+
+ (** freshness **)
+
+ val finite_premss = map (fn aT =>
+ map (fn (f, T) => HOLogic.mk_Trueprop
+ (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
+ (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ f)))
+ (rec_fns ~~ rec_fn_Ts)) dt_atomTs;
+
+ val rec_fns' = map (augment_sort thy11 fs_cp_sort) rec_fns;
+
+ val rec_fresh_thms = map (fn ((aT, eqvt_ths), finite_prems) =>
+ let
+ val name = Long_Name.base_name (fst (dest_Type aT));
+ val fs_name = PureThy.get_thm thy11 ("fs_" ^ name ^ "1");
+ val a = Free ("a", aT);
+ val freshs = map (fn (f, fT) => HOLogic.mk_Trueprop
+ (fresh_const aT fT $ a $ f)) (rec_fns ~~ rec_fn_Ts)
+ in
+ map (fn (((T, U), R), eqvt_th) =>
+ let
+ val x = Free ("x", augment_sort_typ thy11 fs_cp_sort T);
+ val y = Free ("y", U);
+ val y' = Free ("y'", U)
+ in
+ standard (Goal.prove (ProofContext.init thy11) []
+ (map (augment_sort thy11 fs_cp_sort)
+ (finite_prems @
+ [HOLogic.mk_Trueprop (R $ x $ y),
+ HOLogic.mk_Trueprop (HOLogic.mk_all ("y'", U,
+ HOLogic.mk_imp (R $ x $ y', HOLogic.mk_eq (y', y)))),
+ HOLogic.mk_Trueprop (fresh_const aT T $ a $ x)] @
+ freshs))
+ (HOLogic.mk_Trueprop (fresh_const aT U $ a $ y))
+ (fn {prems, context} =>
+ let
+ val (finite_prems, rec_prem :: unique_prem ::
+ fresh_prems) = chop (length finite_prems) prems;
+ val unique_prem' = unique_prem RS spec RS mp;
+ val unique = [unique_prem', unique_prem' RS sym] MRS trans;
+ val _ $ (_ $ (_ $ S $ _)) $ _ = prop_of supports_fresh;
+ val tuple = foldr1 HOLogic.mk_prod (x :: rec_fns')
+ in EVERY
+ [rtac (Drule.cterm_instantiate
+ [(cterm_of thy11 S,
+ cterm_of thy11 (Const ("Nominal.supp",
+ fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))]
+ supports_fresh) 1,
+ simp_tac (HOL_basic_ss addsimps
+ [supports_def, symmetric fresh_def, fresh_prod]) 1,
+ REPEAT_DETERM (resolve_tac [allI, impI] 1),
+ REPEAT_DETERM (etac conjE 1),
+ rtac unique 1,
+ SUBPROOF (fn {prems = prems', params = [a, b], ...} => EVERY
+ [cut_facts_tac [rec_prem] 1,
+ rtac (Thm.instantiate ([],
+ [(cterm_of thy11 (Var (("pi", 0), mk_permT aT)),
+ cterm_of thy11 (perm_of_pair (term_of a, term_of b)))]) eqvt_th) 1,
+ asm_simp_tac (HOL_ss addsimps
+ (prems' @ perm_swap @ perm_fresh_fresh)) 1]) context 1,
+ rtac rec_prem 1,
+ simp_tac (HOL_ss addsimps (fs_name ::
+ supp_prod :: finite_Un :: finite_prems)) 1,
+ simp_tac (HOL_ss addsimps (symmetric fresh_def ::
+ fresh_prod :: fresh_prems)) 1]
+ end))
+ end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ eqvt_ths)
+ end) (dt_atomTs ~~ rec_equiv_thms' ~~ finite_premss);
+
+ (** uniqueness **)
+
+ val fun_tuple = foldr1 HOLogic.mk_prod (rec_ctxt :: rec_fns);
+ val fun_tupleT = fastype_of fun_tuple;
+ val rec_unique_frees =
+ DatatypeProp.indexify_names (replicate (length recTs) "x") ~~ recTs;
+ val rec_unique_frees'' = map (fn (s, T) => (s ^ "'", T)) rec_unique_frees;
+ val rec_unique_frees' =
+ DatatypeProp.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts;
+ val rec_unique_concls = map (fn ((x, U), R) =>
+ Const ("Ex1", (U --> HOLogic.boolT) --> HOLogic.boolT) $
+ Abs ("y", U, R $ Free x $ Bound 0))
+ (rec_unique_frees ~~ rec_result_Ts ~~ rec_sets);
+
+ val induct_aux_rec = Drule.cterm_instantiate
+ (map (pairself (cterm_of thy11) o apsnd (augment_sort thy11 fs_cp_sort))
+ (map (fn (aT, f) => (Logic.varify f, Abs ("z", HOLogic.unitT,
+ Const ("Nominal.supp", fun_tupleT --> HOLogic.mk_setT aT) $ fun_tuple)))
+ fresh_fs @
+ map (fn (((P, T), (x, U)), Q) =>
+ (Var ((P, 0), Logic.varifyT (fsT' --> T --> HOLogic.boolT)),
+ Abs ("z", HOLogic.unitT, absfree (x, U, Q))))
+ (pnames ~~ recTs ~~ rec_unique_frees ~~ rec_unique_concls) @
+ map (fn (s, T) => (Var ((s, 0), Logic.varifyT T), Free (s, T)))
+ rec_unique_frees)) induct_aux;
+
+ fun obtain_fresh_name vs ths rec_fin_supp T (freshs1, freshs2, ctxt) =
+ let
+ val p = foldr1 HOLogic.mk_prod (vs @ freshs1);
+ val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
+ (HOLogic.exists_const T $ Abs ("x", T,
+ fresh_const T (fastype_of p) $ Bound 0 $ p)))
+ (fn _ => EVERY
+ [cut_facts_tac ths 1,
+ REPEAT_DETERM (dresolve_tac (the (AList.lookup op = rec_fin_supp T)) 1),
+ resolve_tac exists_fresh' 1,
+ asm_simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]);
+ val (([cx], ths), ctxt') = Obtain.result
+ (fn _ => EVERY
+ [etac exE 1,
+ full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1,
+ REPEAT (etac conjE 1)])
+ [ex] ctxt
+ in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end;
+
+ val finite_ctxt_prems = map (fn aT =>
+ HOLogic.mk_Trueprop
+ (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
+ (Const ("Nominal.supp", fsT' --> HOLogic.mk_setT aT) $ rec_ctxt))) dt_atomTs;
+
+ val rec_unique_thms = split_conj_thm (Goal.prove
+ (ProofContext.init thy11) (map fst rec_unique_frees)
+ (map (augment_sort thy11 fs_cp_sort)
+ (List.concat finite_premss @ finite_ctxt_prems @ rec_prems @ rec_prems'))
+ (augment_sort thy11 fs_cp_sort
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj rec_unique_concls)))
+ (fn {prems, context} =>
+ let
+ val k = length rec_fns;
+ val (finite_thss, ths1) = fold_map (fn T => fn xs =>
+ apfst (pair T) (chop k xs)) dt_atomTs prems;
+ val (finite_ctxt_ths, ths2) = chop (length dt_atomTs) ths1;
+ val (P_ind_ths, fcbs) = chop k ths2;
+ val P_ths = map (fn th => th RS mp) (split_conj_thm
+ (Goal.prove context
+ (map fst (rec_unique_frees'' @ rec_unique_frees')) []
+ (augment_sort thy11 fs_cp_sort
+ (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+ (map (fn (((x, y), S), P) => HOLogic.mk_imp
+ (S $ Free x $ Free y, P $ (Free y)))
+ (rec_unique_frees'' ~~ rec_unique_frees' ~~
+ rec_sets ~~ rec_preds)))))
+ (fn _ =>
+ rtac rec_induct 1 THEN
+ REPEAT ((resolve_tac P_ind_ths THEN_ALL_NEW assume_tac) 1))));
+ val rec_fin_supp_thms' = map
+ (fn (ths, (T, fin_ths)) => (T, map (curry op MRS fin_ths) ths))
+ (rec_fin_supp_thms ~~ finite_thss);
+ in EVERY
+ ([rtac induct_aux_rec 1] @
+ maps (fn ((_, finite_ths), finite_th) =>
+ [cut_facts_tac (finite_th :: finite_ths) 1,
+ asm_simp_tac (HOL_ss addsimps [supp_prod, finite_Un]) 1])
+ (finite_thss ~~ finite_ctxt_ths) @
+ maps (fn ((_, idxss), elim) => maps (fn idxs =>
+ [full_simp_tac (HOL_ss addsimps [symmetric fresh_def, supp_prod, Un_iff]) 1,
+ REPEAT_DETERM (eresolve_tac [conjE, ex1E] 1),
+ rtac ex1I 1,
+ (resolve_tac rec_intrs THEN_ALL_NEW atac) 1,
+ rotate_tac ~1 1,
+ ((DETERM o etac elim) THEN_ALL_NEW full_simp_tac
+ (HOL_ss addsimps List.concat distinct_thms)) 1] @
+ (if null idxs then [] else [hyp_subst_tac 1,
+ SUBPROOF (fn {asms, concl, prems = prems', params, context = context', ...} =>
+ let
+ val SOME prem = find_first (can (HOLogic.dest_eq o
+ HOLogic.dest_Trueprop o prop_of)) prems';
+ val _ $ (_ $ lhs $ rhs) = prop_of prem;
+ val _ $ (_ $ lhs' $ rhs') = term_of concl;
+ val rT = fastype_of lhs';
+ val (c, cargsl) = strip_comb lhs;
+ val cargsl' = partition_cargs idxs cargsl;
+ val boundsl = List.concat (map fst cargsl');
+ val (_, cargsr) = strip_comb rhs;
+ val cargsr' = partition_cargs idxs cargsr;
+ val boundsr = List.concat (map fst cargsr');
+ val (params1, _ :: params2) =
+ chop (length params div 2) (map term_of params);
+ val params' = params1 @ params2;
+ val rec_prems = filter (fn th => case prop_of th of
+ _ $ p => (case head_of p of
+ Const (s, _) => s mem rec_set_names
+ | _ => false)
+ | _ => false) prems';
+ val fresh_prems = filter (fn th => case prop_of th of
+ _ $ (Const ("Nominal.fresh", _) $ _ $ _) => true
+ | _ $ (Const ("Not", _) $ _) => true
+ | _ => false) prems';
+ val Ts = map fastype_of boundsl;
+
+ val _ = warning "step 1: obtaining fresh names";
+ val (freshs1, freshs2, context'') = fold
+ (obtain_fresh_name (rec_ctxt :: rec_fns' @ params')
+ (List.concat (map snd finite_thss) @
+ finite_ctxt_ths @ rec_prems)
+ rec_fin_supp_thms')
+ Ts ([], [], context');
+ val pi1 = map perm_of_pair (boundsl ~~ freshs1);
+ val rpi1 = rev pi1;
+ val pi2 = map perm_of_pair (boundsr ~~ freshs1);
+ val rpi2 = rev pi2;
+
+ val fresh_prems' = mk_not_sym fresh_prems;
+ val freshs2' = mk_not_sym freshs2;
+
+ (** as, bs, cs # K as ts, K bs us **)
+ val _ = warning "step 2: as, bs, cs # K as ts, K bs us";
+ val prove_fresh_ss = HOL_ss addsimps
+ (finite_Diff :: List.concat fresh_thms @
+ fs_atoms @ abs_fresh @ abs_supp @ fresh_atm);
+ (* FIXME: avoid asm_full_simp_tac ? *)
+ fun prove_fresh ths y x = Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (fresh_const
+ (fastype_of x) (fastype_of y) $ x $ y))
+ (fn _ => cut_facts_tac ths 1 THEN asm_full_simp_tac prove_fresh_ss 1);
+ val constr_fresh_thms =
+ map (prove_fresh fresh_prems lhs) boundsl @
+ map (prove_fresh fresh_prems rhs) boundsr @
+ map (prove_fresh freshs2 lhs) freshs1 @
+ map (prove_fresh freshs2 rhs) freshs1;
+
+ (** pi1 o (K as ts) = pi2 o (K bs us) **)
+ val _ = warning "step 3: pi1 o (K as ts) = pi2 o (K bs us)";
+ val pi1_pi2_eq = Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (fold_rev (mk_perm []) pi1 lhs, fold_rev (mk_perm []) pi2 rhs)))
+ (fn _ => EVERY
+ [cut_facts_tac constr_fresh_thms 1,
+ asm_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh) 1,
+ rtac prem 1]);
+
+ (** pi1 o ts = pi2 o us **)
+ val _ = warning "step 4: pi1 o ts = pi2 o us";
+ val pi1_pi2_eqs = map (fn (t, u) =>
+ Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (fold_rev (mk_perm []) pi1 t, fold_rev (mk_perm []) pi2 u)))
+ (fn _ => EVERY
+ [cut_facts_tac [pi1_pi2_eq] 1,
+ asm_full_simp_tac (HOL_ss addsimps
+ (calc_atm @ List.concat perm_simps' @
+ fresh_prems' @ freshs2' @ abs_perm @
+ alpha @ List.concat inject_thms)) 1]))
+ (map snd cargsl' ~~ map snd cargsr');
+
+ (** pi1^-1 o pi2 o us = ts **)
+ val _ = warning "step 5: pi1^-1 o pi2 o us = ts";
+ val rpi1_pi2_eqs = map (fn ((t, u), eq) =>
+ Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (fold_rev (mk_perm []) (rpi1 @ pi2) u, t)))
+ (fn _ => simp_tac (HOL_ss addsimps
+ ((eq RS sym) :: perm_swap)) 1))
+ (map snd cargsl' ~~ map snd cargsr' ~~ pi1_pi2_eqs);
+
+ val (rec_prems1, rec_prems2) =
+ chop (length rec_prems div 2) rec_prems;
+
+ (** (ts, pi1^-1 o pi2 o vs) in rec_set **)
+ val _ = warning "step 6: (ts, pi1^-1 o pi2 o vs) in rec_set";
+ val rec_prems' = map (fn th =>
+ let
+ val _ $ (S $ x $ y) = prop_of th;
+ val Const (s, _) = head_of S;
+ val k = find_index (equal s) rec_set_names;
+ val pi = rpi1 @ pi2;
+ fun mk_pi z = fold_rev (mk_perm []) pi z;
+ fun eqvt_tac p =
+ let
+ val U as Type (_, [Type (_, [T, _])]) = fastype_of p;
+ val l = find_index (equal T) dt_atomTs;
+ val th = List.nth (List.nth (rec_equiv_thms', l), k);
+ val th' = Thm.instantiate ([],
+ [(cterm_of thy11 (Var (("pi", 0), U)),
+ cterm_of thy11 p)]) th;
+ in rtac th' 1 end;
+ val th' = Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (S $ mk_pi x $ mk_pi y))
+ (fn _ => EVERY
+ (map eqvt_tac pi @
+ [simp_tac (HOL_ss addsimps (fresh_prems' @ freshs2' @
+ perm_swap @ perm_fresh_fresh)) 1,
+ rtac th 1]))
+ in
+ Simplifier.simplify
+ (HOL_basic_ss addsimps rpi1_pi2_eqs) th'
+ end) rec_prems2;
+
+ val ihs = filter (fn th => case prop_of th of
+ _ $ (Const ("All", _) $ _) => true | _ => false) prems';
+
+ (** pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs **)
+ val _ = warning "step 7: pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs";
+ val rec_eqns = map (fn (th, ih) =>
+ let
+ val th' = th RS (ih RS spec RS mp) RS sym;
+ val _ $ (_ $ lhs $ rhs) = prop_of th';
+ fun strip_perm (_ $ _ $ t) = strip_perm t
+ | strip_perm t = t;
+ in
+ Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (fold_rev (mk_perm []) pi1 lhs,
+ fold_rev (mk_perm []) pi2 (strip_perm rhs))))
+ (fn _ => simp_tac (HOL_basic_ss addsimps
+ (th' :: perm_swap)) 1)
+ end) (rec_prems' ~~ ihs);
+
+ (** as # rs **)
+ val _ = warning "step 8: as # rs";
+ val rec_freshs = List.concat
+ (map (fn (rec_prem, ih) =>
+ let
+ val _ $ (S $ x $ (y as Free (_, T))) =
+ prop_of rec_prem;
+ val k = find_index (equal S) rec_sets;
+ val atoms = List.concat (List.mapPartial (fn (bs, z) =>
+ if z = x then NONE else SOME bs) cargsl')
+ in
+ map (fn a as Free (_, aT) =>
+ let val l = find_index (equal aT) dt_atomTs;
+ in
+ Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (fresh_const aT T $ a $ y))
+ (fn _ => EVERY
+ (rtac (List.nth (List.nth (rec_fresh_thms, l), k)) 1 ::
+ map (fn th => rtac th 1)
+ (snd (List.nth (finite_thss, l))) @
+ [rtac rec_prem 1, rtac ih 1,
+ REPEAT_DETERM (resolve_tac fresh_prems 1)]))
+ end) atoms
+ end) (rec_prems1 ~~ ihs));
+
+ (** as # fK as ts rs , bs # fK bs us vs **)
+ val _ = warning "step 9: as # fK as ts rs , bs # fK bs us vs";
+ fun prove_fresh_result (a as Free (_, aT)) =
+ Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ rhs'))
+ (fn _ => EVERY
+ [resolve_tac fcbs 1,
+ REPEAT_DETERM (resolve_tac
+ (fresh_prems @ rec_freshs) 1),
+ REPEAT_DETERM (resolve_tac (maps snd rec_fin_supp_thms') 1
+ THEN resolve_tac rec_prems 1),
+ resolve_tac P_ind_ths 1,
+ REPEAT_DETERM (resolve_tac (P_ths @ rec_prems) 1)]);
+
+ val fresh_results'' = map prove_fresh_result boundsl;
+
+ fun prove_fresh_result'' ((a as Free (_, aT), b), th) =
+ let val th' = Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (fresh_const aT rT $
+ fold_rev (mk_perm []) (rpi2 @ pi1) a $
+ fold_rev (mk_perm []) (rpi2 @ pi1) rhs'))
+ (fn _ => simp_tac (HOL_ss addsimps fresh_bij) 1 THEN
+ rtac th 1)
+ in
+ Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (fresh_const aT rT $ b $ lhs'))
+ (fn _ => EVERY
+ [cut_facts_tac [th'] 1,
+ full_simp_tac (Simplifier.theory_context thy11 HOL_ss
+ addsimps rec_eqns @ pi1_pi2_eqs @ perm_swap
+ addsimprocs [NominalPermeq.perm_simproc_app]) 1,
+ full_simp_tac (HOL_ss addsimps (calc_atm @
+ fresh_prems' @ freshs2' @ perm_fresh_fresh)) 1])
+ end;
+
+ val fresh_results = fresh_results'' @ map prove_fresh_result''
+ (boundsl ~~ boundsr ~~ fresh_results'');
+
+ (** cs # fK as ts rs , cs # fK bs us vs **)
+ val _ = warning "step 10: cs # fK as ts rs , cs # fK bs us vs";
+ fun prove_fresh_result' recs t (a as Free (_, aT)) =
+ Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ t))
+ (fn _ => EVERY
+ [cut_facts_tac recs 1,
+ REPEAT_DETERM (dresolve_tac
+ (the (AList.lookup op = rec_fin_supp_thms' aT)) 1),
+ NominalPermeq.fresh_guess_tac
+ (HOL_ss addsimps (freshs2 @
+ fs_atoms @ fresh_atm @
+ List.concat (map snd finite_thss))) 1]);
+
+ val fresh_results' =
+ map (prove_fresh_result' rec_prems1 rhs') freshs1 @
+ map (prove_fresh_result' rec_prems2 lhs') freshs1;
+
+ (** pi1 o (fK as ts rs) = pi2 o (fK bs us vs) **)
+ val _ = warning "step 11: pi1 o (fK as ts rs) = pi2 o (fK bs us vs)";
+ val pi1_pi2_result = Goal.prove context'' [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (fold_rev (mk_perm []) pi1 rhs', fold_rev (mk_perm []) pi2 lhs')))
+ (fn _ => simp_tac (Simplifier.context context'' HOL_ss
+ addsimps pi1_pi2_eqs @ rec_eqns
+ addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
+ TRY (simp_tac (HOL_ss addsimps
+ (fresh_prems' @ freshs2' @ calc_atm @ perm_fresh_fresh)) 1));
+
+ val _ = warning "final result";
+ val final = Goal.prove context'' [] [] (term_of concl)
+ (fn _ => cut_facts_tac [pi1_pi2_result RS sym] 1 THEN
+ full_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh @
+ fresh_results @ fresh_results') 1);
+ val final' = ProofContext.export context'' context' [final];
+ val _ = warning "finished!"
+ in
+ resolve_tac final' 1
+ end) context 1])) idxss) (ndescr ~~ rec_elims))
+ end));
+
+ val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
+
+ (* define primrec combinators *)
+
+ val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
+ val reccomb_names = map (Sign.full_bname thy11)
+ (if length descr'' = 1 then [big_reccomb_name] else
+ (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
+ (1 upto (length descr''))));
+ val reccombs = map (fn ((name, T), T') => list_comb
+ (Const (name, rec_fn_Ts @ [T] ---> T'), rec_fns))
+ (reccomb_names ~~ recTs ~~ rec_result_Ts);
+
+ val (reccomb_defs, thy12) =
+ thy11
+ |> Sign.add_consts_i (map (fn ((name, T), T') =>
+ (Binding.name (Long_Name.base_name name), rec_fn_Ts @ [T] ---> T', NoSyn))
+ (reccomb_names ~~ recTs ~~ rec_result_Ts))
+ |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
+ (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T,
+ Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
+ set $ Free ("x", T) $ Free ("y", T'))))))
+ (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts));
+
+ (* prove characteristic equations for primrec combinators *)
+
+ val rec_thms = map (fn (prems, concl) =>
+ let
+ val _ $ (_ $ (_ $ x) $ _) = concl;
+ val (_, cargs) = strip_comb x;
+ val ps = map (fn (x as Free (_, T), i) =>
+ (Free ("x" ^ string_of_int i, T), x)) (cargs ~~ (1 upto length cargs));
+ val concl' = subst_atomic_types (rec_result_Ts' ~~ rec_result_Ts) concl;
+ val prems' = List.concat finite_premss @ finite_ctxt_prems @
+ rec_prems @ rec_prems' @ map (subst_atomic ps) prems;
+ fun solve rules prems = resolve_tac rules THEN_ALL_NEW
+ (resolve_tac prems THEN_ALL_NEW atac)
+ in
+ Goal.prove_global thy12 []
+ (map (augment_sort thy12 fs_cp_sort) prems')
+ (augment_sort thy12 fs_cp_sort concl')
+ (fn {prems, ...} => EVERY
+ [rewrite_goals_tac reccomb_defs,
+ rtac the1_equality 1,
+ solve rec_unique_thms prems 1,
+ resolve_tac rec_intrs 1,
+ REPEAT (solve (prems @ rec_total_thms) prems 1)])
+ end) (rec_eq_prems ~~
+ DatatypeProp.make_primrecs new_type_names descr' sorts thy12);
+
+ val dt_infos = map (make_dt_info pdescr sorts induct reccomb_names rec_thms)
+ ((0 upto length descr1 - 1) ~~ descr1 ~~ distinct_thms ~~ inject_thms);
+
+ (* FIXME: theorems are stored in database for testing only *)
+ val (_, thy13) = thy12 |>
+ PureThy.add_thmss
+ [((Binding.name "rec_equiv", List.concat rec_equiv_thms), []),
+ ((Binding.name "rec_equiv'", List.concat rec_equiv_thms'), []),
+ ((Binding.name "rec_fin_supp", List.concat rec_fin_supp_thms), []),
+ ((Binding.name "rec_fresh", List.concat rec_fresh_thms), []),
+ ((Binding.name "rec_unique", map standard rec_unique_thms), []),
+ ((Binding.name "recs", rec_thms), [])] ||>
+ Sign.parent_path ||>
+ map_nominal_datatypes (fold Symtab.update dt_infos);
+
+ in
+ thy13
+ end;
+
+val add_nominal_datatype = gen_add_nominal_datatype Datatype.read_typ;
+
+
+(* FIXME: The following stuff should be exported by Datatype *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val datatype_decl =
+ Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix --
+ (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix));
+
+fun mk_datatype args =
+ let
+ val names = map (fn ((((NONE, _), t), _), _) => t | ((((SOME t, _), _), _), _) => t) args;
+ val specs = map (fn ((((_, vs), t), mx), cons) =>
+ (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
+ in add_nominal_datatype DatatypeAux.default_datatype_config names specs end;
+
+val _ =
+ OuterSyntax.command "nominal_datatype" "define inductive datatypes" K.thy_decl
+ (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
+
+end;
+
+end
--- a/src/HOL/Nominal/nominal_atoms.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Nominal/nominal_atoms.ML Fri Jun 19 17:23:21 2009 +0200
@@ -101,7 +101,7 @@
val (_,thy1) =
fold_map (fn ak => fn thy =>
let val dt = ([], Binding.name ak, NoSyn, [(Binding.name ak, [@{typ nat}], NoSyn)])
- val ({inject,case_thms,...},thy1) = DatatypePackage.add_datatype
+ val ({inject,case_thms,...},thy1) = Datatype.add_datatype
DatatypeAux.default_datatype_config [ak] [dt] thy
val inject_flat = flat inject
val ak_type = Type (Sign.intern_type thy1 ak,[])
@@ -191,7 +191,7 @@
thy |> Sign.add_consts_i [(Binding.name ("swap_" ^ ak_name), swapT, NoSyn)]
|> PureThy.add_defs_unchecked true [((Binding.name name, def2),[])]
|> snd
- |> OldPrimrecPackage.add_primrec_unchecked_i "" [(("", def1),[])]
+ |> OldPrimrec.add_primrec_unchecked_i "" [(("", def1),[])]
end) ak_names_types thy1;
(* declares a permutation function for every atom-kind acting *)
@@ -219,7 +219,7 @@
Const (swap_name, swapT) $ x $ (Const (qu_prm_name, prmT) $ xs $ a)));
in
thy |> Sign.add_consts_i [(Binding.name prm_name, mk_permT T --> T --> T, NoSyn)]
- |> OldPrimrecPackage.add_primrec_unchecked_i "" [(("", def1), []),(("", def2), [])]
+ |> OldPrimrec.add_primrec_unchecked_i "" [(("", def1), []),(("", def2), [])]
end) ak_names_types thy3;
(* defines permutation functions for all combinations of atom-kinds; *)
--- a/src/HOL/Nominal/nominal_inductive.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Nominal/nominal_inductive.ML Fri Jun 19 17:23:21 2009 +0200
@@ -53,7 +53,7 @@
fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of
(Const (s, T), ts) => (case strip_type T of
(Ts, Type (tname, _)) =>
- (case NominalPackage.get_nominal_datatype thy tname of
+ (case Nominal.get_nominal_datatype thy tname of
NONE => fold (add_binders thy i) ts bs
| SOME {descr, index, ...} => (case AList.lookup op =
(#3 (the (AList.lookup op = descr index))) s of
@@ -148,11 +148,11 @@
let
val thy = ProofContext.theory_of ctxt;
val ({names, ...}, {raw_induct, intrs, elims, ...}) =
- InductivePackage.the_inductive ctxt (Sign.intern_const thy s);
- val ind_params = InductivePackage.params_of raw_induct;
+ Inductive.the_inductive ctxt (Sign.intern_const thy s);
+ val ind_params = Inductive.params_of raw_induct;
val raw_induct = atomize_induct ctxt raw_induct;
val elims = map (atomize_induct ctxt) elims;
- val monos = InductivePackage.get_monos ctxt;
+ val monos = Inductive.get_monos ctxt;
val eqvt_thms = NominalThmDecls.get_eqvt_thms ctxt;
val _ = (case names \\ fold (Term.add_const_names o Thm.prop_of) eqvt_thms [] of
[] => ()
@@ -230,7 +230,7 @@
else NONE) xs @ mk_distinct xs;
fun mk_fresh (x, T) = HOLogic.mk_Trueprop
- (NominalPackage.fresh_const T fsT $ x $ Bound 0);
+ (Nominal.fresh_const T fsT $ x $ Bound 0);
val (prems', prems'') = split_list (map (fn (params, bvars, prems, (p, ts)) =>
let
@@ -254,7 +254,7 @@
val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
(map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem,
HOLogic.list_all (ind_vars, lift_pred p
- (map (fold_rev (NominalPackage.mk_perm ind_Ts)
+ (map (fold_rev (Nominal.mk_perm ind_Ts)
(map Bound (length atomTs downto 1))) ts)))) concls));
val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
@@ -268,7 +268,7 @@
else map_term (split_conj (K o I) names) prem prem) prems, q))))
(mk_distinct bvars @
maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop
- (NominalPackage.fresh_const U T $ u $ t)) bvars)
+ (Nominal.fresh_const U T $ u $ t)) bvars)
(ts ~~ binder_types (fastype_of p)))) prems;
val perm_pi_simp = PureThy.get_thms thy "perm_pi_simp";
@@ -296,7 +296,7 @@
val p = foldr1 HOLogic.mk_prod (map protect ts @ freshs1);
val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
(HOLogic.exists_const T $ Abs ("x", T,
- NominalPackage.fresh_const T (fastype_of p) $
+ Nominal.fresh_const T (fastype_of p) $
Bound 0 $ p)))
(fn _ => EVERY
[resolve_tac exists_fresh' 1,
@@ -325,13 +325,13 @@
(fn (Bound i, T) => (nth params' (length params' - i), T)
| (t, T) => (t, T)) bvars;
val pi_bvars = map (fn (t, _) =>
- fold_rev (NominalPackage.mk_perm []) pis t) bvars';
+ fold_rev (Nominal.mk_perm []) pis t) bvars';
val (P, ts) = strip_comb (HOLogic.dest_Trueprop (term_of concl));
val (freshs1, freshs2, ctxt'') = fold
(obtain_fresh_name (ts @ pi_bvars))
(map snd bvars') ([], [], ctxt');
- val freshs2' = NominalPackage.mk_not_sym freshs2;
- val pis' = map NominalPackage.perm_of_pair (pi_bvars ~~ freshs1);
+ val freshs2' = Nominal.mk_not_sym freshs2;
+ val pis' = map Nominal.perm_of_pair (pi_bvars ~~ freshs1);
fun concat_perm pi1 pi2 =
let val T = fastype_of pi1
in if T = fastype_of pi2 then
@@ -343,11 +343,11 @@
(Vartab.empty, Vartab.empty);
val ihyp' = Thm.instantiate ([], map (pairself (cterm_of thy))
(map (Envir.subst_vars env) vs ~~
- map (fold_rev (NominalPackage.mk_perm [])
+ map (fold_rev (Nominal.mk_perm [])
(rev pis' @ pis)) params' @ [z])) ihyp;
fun mk_pi th =
Simplifier.simplify (HOL_basic_ss addsimps [@{thm id_apply}]
- addsimprocs [NominalPackage.perm_simproc])
+ addsimprocs [Nominal.perm_simproc])
(Simplifier.simplify eqvt_ss
(fold_rev (mk_perm_bool o cterm_of thy)
(rev pis' @ pis) th));
@@ -369,13 +369,13 @@
| _ $ (_ $ (_ $ lhs $ rhs)) =>
(curry (HOLogic.mk_not o HOLogic.mk_eq), lhs, rhs));
val th'' = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop
- (bop (fold_rev (NominalPackage.mk_perm []) pis lhs)
- (fold_rev (NominalPackage.mk_perm []) pis rhs)))
+ (bop (fold_rev (Nominal.mk_perm []) pis lhs)
+ (fold_rev (Nominal.mk_perm []) pis rhs)))
(fn _ => simp_tac (HOL_basic_ss addsimps
(fresh_bij @ perm_bij)) 1 THEN rtac th' 1)
in Simplifier.simplify (eqvt_ss addsimps fresh_atm) th'' end)
vc_compat_ths;
- val vc_compat_ths'' = NominalPackage.mk_not_sym vc_compat_ths';
+ val vc_compat_ths'' = Nominal.mk_not_sym vc_compat_ths';
(** Since swap_simps simplifies (pi :: 'a prm) o (x :: 'b) to x **)
(** we have to pre-simplify the rewrite rules **)
val swap_simps_ss = HOL_ss addsimps swap_simps @
@@ -383,14 +383,14 @@
(vc_compat_ths'' @ freshs2');
val th = Goal.prove ctxt'' [] []
(HOLogic.mk_Trueprop (list_comb (P $ hd ts,
- map (fold (NominalPackage.mk_perm []) pis') (tl ts))))
+ map (fold (Nominal.mk_perm []) pis') (tl ts))))
(fn _ => EVERY ([simp_tac eqvt_ss 1, rtac ihyp' 1,
REPEAT_DETERM_N (nprems_of ihyp - length gprems)
(simp_tac swap_simps_ss 1),
REPEAT_DETERM_N (length gprems)
(simp_tac (HOL_basic_ss
addsimps [inductive_forall_def']
- addsimprocs [NominalPackage.perm_simproc]) 1 THEN
+ addsimprocs [Nominal.perm_simproc]) 1 THEN
resolve_tac gprems2 1)]));
val final = Goal.prove ctxt'' [] [] (term_of concl)
(fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (HOL_ss
@@ -435,7 +435,7 @@
((ps, qs, is, map (curry subst_bounds (rev ts)) prems), ctxt')
end) (prems ~~ avoids) ctxt')
end)
- (InductivePackage.partition_rules' raw_induct (intrs ~~ avoids') ~~
+ (Inductive.partition_rules' raw_induct (intrs ~~ avoids') ~~
elims);
val cases_prems' =
@@ -448,7 +448,7 @@
(Logic.list_implies
(mk_distinct qs @
maps (fn (t, T) => map (fn u => HOLogic.mk_Trueprop
- (NominalPackage.fresh_const T (fastype_of u) $ t $ u))
+ (Nominal.fresh_const T (fastype_of u) $ t $ u))
args) qs,
HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
(map HOLogic.dest_Trueprop prems))),
@@ -499,13 +499,13 @@
chop (length vc_compat_ths - length args * length qs)
(map (first_order_mrs hyps2) vc_compat_ths);
val vc_compat_ths' =
- NominalPackage.mk_not_sym vc_compat_ths1 @
+ Nominal.mk_not_sym vc_compat_ths1 @
flat (fst (fold_map inst_fresh hyps1 vc_compat_ths2));
val (freshs1, freshs2, ctxt3) = fold
(obtain_fresh_name (args @ map fst qs @ params'))
(map snd qs) ([], [], ctxt2);
- val freshs2' = NominalPackage.mk_not_sym freshs2;
- val pis = map (NominalPackage.perm_of_pair)
+ val freshs2' = Nominal.mk_not_sym freshs2;
+ val pis = map (Nominal.perm_of_pair)
((freshs1 ~~ map fst qs) @ (params' ~~ freshs1));
val mk_pis = fold_rev mk_perm_bool (map (cterm_of thy) pis);
val obj = cterm_of thy (foldr1 HOLogic.mk_conj (map (map_aterms
@@ -513,7 +513,7 @@
if x mem args then x
else (case AList.lookup op = tab x of
SOME y => y
- | NONE => fold_rev (NominalPackage.mk_perm []) pis x)
+ | NONE => fold_rev (Nominal.mk_perm []) pis x)
| x => x) o HOLogic.dest_Trueprop o prop_of) case_hyps));
val inst = Thm.first_order_match (Thm.dest_arg
(Drule.strip_imp_concl (hd (cprems_of case_hyp))), obj);
@@ -522,7 +522,7 @@
rtac (Thm.instantiate inst case_hyp) 1 THEN
SUBPROOF (fn {prems = fresh_hyps, ...} =>
let
- val fresh_hyps' = NominalPackage.mk_not_sym fresh_hyps;
+ val fresh_hyps' = Nominal.mk_not_sym fresh_hyps;
val case_ss = cases_eqvt_ss addsimps freshs2' @
simp_fresh_atm (vc_compat_ths' @ fresh_hyps');
val fresh_fresh_ss = case_ss addsimps perm_fresh_fresh;
@@ -548,13 +548,13 @@
val rec_name = space_implode "_" (map Long_Name.base_name names);
val rec_qualified = Binding.qualify false rec_name;
val ind_case_names = RuleCases.case_names induct_cases;
- val induct_cases' = InductivePackage.partition_rules' raw_induct
+ val induct_cases' = Inductive.partition_rules' raw_induct
(intrs ~~ induct_cases);
val thss' = map (map atomize_intr) thss;
- val thsss = InductivePackage.partition_rules' raw_induct (intrs ~~ thss');
+ val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss');
val strong_raw_induct =
- mk_ind_proof ctxt thss' |> InductivePackage.rulify;
- val strong_cases = map (mk_cases_proof ##> InductivePackage.rulify)
+ mk_ind_proof ctxt thss' |> Inductive.rulify;
+ val strong_cases = map (mk_cases_proof ##> Inductive.rulify)
(thsss ~~ elims ~~ cases_prems ~~ cases_prems');
val strong_induct =
if length names > 1 then
@@ -587,17 +587,17 @@
let
val thy = ProofContext.theory_of ctxt;
val ({names, ...}, {raw_induct, intrs, elims, ...}) =
- InductivePackage.the_inductive ctxt (Sign.intern_const thy s);
+ Inductive.the_inductive ctxt (Sign.intern_const thy s);
val raw_induct = atomize_induct ctxt raw_induct;
val elims = map (atomize_induct ctxt) elims;
val intrs = map atomize_intr intrs;
- val monos = InductivePackage.get_monos ctxt;
- val intrs' = InductivePackage.unpartition_rules intrs
+ val monos = Inductive.get_monos ctxt;
+ val intrs' = Inductive.unpartition_rules intrs
(map (fn (((s, ths), (_, k)), th) =>
- (s, ths ~~ InductivePackage.infer_intro_vars th k ths))
- (InductivePackage.partition_rules raw_induct intrs ~~
- InductivePackage.arities_of raw_induct ~~ elims));
- val k = length (InductivePackage.params_of raw_induct);
+ (s, ths ~~ Inductive.infer_intro_vars th k ths))
+ (Inductive.partition_rules raw_induct intrs ~~
+ Inductive.arities_of raw_induct ~~ elims));
+ val k = length (Inductive.params_of raw_induct);
val atoms' = NominalAtoms.atoms_of thy;
val atoms =
if null xatoms then atoms' else
@@ -635,7 +635,7 @@
val prems'' = map (fn th => Simplifier.simplify eqvt_ss
(mk_perm_bool (cterm_of thy pi) th)) prems';
val intr' = Drule.cterm_instantiate (map (cterm_of thy) vs ~~
- map (cterm_of thy o NominalPackage.mk_perm [] pi o term_of) params)
+ map (cterm_of thy o Nominal.mk_perm [] pi o term_of) params)
intr
in (rtac intr' THEN_ALL_NEW (TRY o resolve_tac prems'')) 1
end) ctxt' 1 st
@@ -655,7 +655,7 @@
val (ts1, ts2) = chop k ts
in
HOLogic.mk_imp (p, list_comb (h, ts1 @
- map (NominalPackage.mk_perm [] pi') ts2))
+ map (Nominal.mk_perm [] pi') ts2))
end) ps)))
(fn {context, ...} => EVERY (rtac raw_induct 1 :: map (fn intr_vs =>
full_simp_tac eqvt_ss 1 THEN
--- a/src/HOL/Nominal/nominal_inductive2.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Nominal/nominal_inductive2.ML Fri Jun 19 17:23:21 2009 +0200
@@ -56,7 +56,7 @@
fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of
(Const (s, T), ts) => (case strip_type T of
(Ts, Type (tname, _)) =>
- (case NominalPackage.get_nominal_datatype thy tname of
+ (case Nominal.get_nominal_datatype thy tname of
NONE => fold (add_binders thy i) ts bs
| SOME {descr, index, ...} => (case AList.lookup op =
(#3 (the (AList.lookup op = descr index))) s of
@@ -154,11 +154,11 @@
let
val thy = ProofContext.theory_of ctxt;
val ({names, ...}, {raw_induct, intrs, elims, ...}) =
- InductivePackage.the_inductive ctxt (Sign.intern_const thy s);
- val ind_params = InductivePackage.params_of raw_induct;
+ Inductive.the_inductive ctxt (Sign.intern_const thy s);
+ val ind_params = Inductive.params_of raw_induct;
val raw_induct = atomize_induct ctxt raw_induct;
val elims = map (atomize_induct ctxt) elims;
- val monos = InductivePackage.get_monos ctxt;
+ val monos = Inductive.get_monos ctxt;
val eqvt_thms = NominalThmDecls.get_eqvt_thms ctxt;
val _ = (case names \\ fold (Term.add_const_names o Thm.prop_of) eqvt_thms [] of
[] => ()
@@ -249,7 +249,7 @@
| lift_prem t = t;
fun mk_fresh (x, T) = HOLogic.mk_Trueprop
- (NominalPackage.fresh_star_const T fsT $ x $ Bound 0);
+ (Nominal.fresh_star_const T fsT $ x $ Bound 0);
val (prems', prems'') = split_list (map (fn (params, sets, prems, (p, ts)) =>
let
@@ -270,7 +270,7 @@
val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
(map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem,
HOLogic.list_all (ind_vars, lift_pred p
- (map (fold_rev (NominalPackage.mk_perm ind_Ts)
+ (map (fold_rev (Nominal.mk_perm ind_Ts)
(map Bound (length atomTs downto 1))) ts)))) concls));
val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
@@ -283,7 +283,7 @@
if null (preds_of ps prem) then SOME prem
else map_term (split_conj (K o I) names) prem prem) prems, q))))
(maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop
- (NominalPackage.fresh_star_const U T $ u $ t)) sets)
+ (Nominal.fresh_star_const U T $ u $ t)) sets)
(ts ~~ binder_types (fastype_of p)) @
map (fn (u, U) => HOLogic.mk_Trueprop (Const (@{const_name finite},
HOLogic.mk_setT U --> HOLogic.boolT) $ u)) sets) |>
@@ -339,7 +339,7 @@
val th2' =
Goal.prove ctxt [] []
(list_all (map (pair "pi") pTs, HOLogic.mk_Trueprop
- (f $ fold_rev (NominalPackage.mk_perm (rev pTs))
+ (f $ fold_rev (Nominal.mk_perm (rev pTs))
(pis1 @ pi :: pis2) l $ r)))
(fn _ => cut_facts_tac [th2] 1 THEN
full_simp_tac (HOL_basic_ss addsimps perm_set_forget) 1) |>
@@ -364,7 +364,7 @@
val params' = map term_of cparams'
val sets' = map (apfst (curry subst_bounds (rev params'))) sets;
val pi_sets = map (fn (t, _) =>
- fold_rev (NominalPackage.mk_perm []) pis t) sets';
+ fold_rev (Nominal.mk_perm []) pis t) sets';
val (P, ts) = strip_comb (HOLogic.dest_Trueprop (term_of concl));
val gprems1 = List.mapPartial (fn (th, t) =>
if null (preds_of ps t) then SOME th
@@ -380,7 +380,7 @@
in
Goal.prove ctxt' [] []
(HOLogic.mk_Trueprop (list_comb (h,
- map (fold_rev (NominalPackage.mk_perm []) pis) ts)))
+ map (fold_rev (Nominal.mk_perm []) pis) ts)))
(fn _ => simp_tac (HOL_basic_ss addsimps
(fresh_star_bij @ finite_ineq)) 1 THEN rtac th' 1)
end) vc_compat_ths vc_compat_vs;
@@ -400,11 +400,11 @@
end;
val pis'' = fold_rev (concat_perm #> map) pis' pis;
val ihyp' = inst_params thy vs_ihypt ihyp
- (map (fold_rev (NominalPackage.mk_perm [])
+ (map (fold_rev (Nominal.mk_perm [])
(pis' @ pis) #> cterm_of thy) params' @ [cterm_of thy z]);
fun mk_pi th =
Simplifier.simplify (HOL_basic_ss addsimps [@{thm id_apply}]
- addsimprocs [NominalPackage.perm_simproc])
+ addsimprocs [Nominal.perm_simproc])
(Simplifier.simplify eqvt_ss
(fold_rev (mk_perm_bool o cterm_of thy)
(pis' @ pis) th));
@@ -419,13 +419,13 @@
(fresh_ths2 ~~ sets);
val th = Goal.prove ctxt'' [] []
(HOLogic.mk_Trueprop (list_comb (P $ hd ts,
- map (fold_rev (NominalPackage.mk_perm []) pis') (tl ts))))
+ map (fold_rev (Nominal.mk_perm []) pis') (tl ts))))
(fn _ => EVERY ([simp_tac eqvt_ss 1, rtac ihyp' 1] @
map (fn th => rtac th 1) fresh_ths3 @
[REPEAT_DETERM_N (length gprems)
(simp_tac (HOL_basic_ss
addsimps [inductive_forall_def']
- addsimprocs [NominalPackage.perm_simproc]) 1 THEN
+ addsimprocs [Nominal.perm_simproc]) 1 THEN
resolve_tac gprems2 1)]));
val final = Goal.prove ctxt'' [] [] (term_of concl)
(fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (HOL_ss
@@ -450,12 +450,12 @@
val rec_name = space_implode "_" (map Long_Name.base_name names);
val rec_qualified = Binding.qualify false rec_name;
val ind_case_names = RuleCases.case_names induct_cases;
- val induct_cases' = InductivePackage.partition_rules' raw_induct
+ val induct_cases' = Inductive.partition_rules' raw_induct
(intrs ~~ induct_cases);
val thss' = map (map atomize_intr) thss;
- val thsss = InductivePackage.partition_rules' raw_induct (intrs ~~ thss');
+ val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss');
val strong_raw_induct =
- mk_ind_proof ctxt thss' |> InductivePackage.rulify;
+ mk_ind_proof ctxt thss' |> Inductive.rulify;
val strong_induct =
if length names > 1 then
(strong_raw_induct, [ind_case_names, RuleCases.consumes 0])
--- a/src/HOL/Nominal/nominal_package.ML Thu Jun 18 18:31:14 2009 -0700
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,2095 +0,0 @@
-(* Title: HOL/Nominal/nominal_package.ML
- Author: Stefan Berghofer and Christian Urban, TU Muenchen
-
-Nominal datatype package for Isabelle/HOL.
-*)
-
-signature NOMINAL_PACKAGE =
-sig
- val add_nominal_datatype : DatatypeAux.datatype_config -> string list ->
- (string list * bstring * mixfix *
- (bstring * string list * mixfix) list) list -> theory -> theory
- type descr
- type nominal_datatype_info
- val get_nominal_datatypes : theory -> nominal_datatype_info Symtab.table
- val get_nominal_datatype : theory -> string -> nominal_datatype_info option
- val mk_perm: typ list -> term -> term -> term
- val perm_of_pair: term * term -> term
- val mk_not_sym: thm list -> thm list
- val perm_simproc: simproc
- val fresh_const: typ -> typ -> term
- val fresh_star_const: typ -> typ -> term
-end
-
-structure NominalPackage : NOMINAL_PACKAGE =
-struct
-
-val finite_emptyI = thm "finite.emptyI";
-val finite_Diff = thm "finite_Diff";
-val finite_Un = thm "finite_Un";
-val Un_iff = thm "Un_iff";
-val In0_eq = thm "In0_eq";
-val In1_eq = thm "In1_eq";
-val In0_not_In1 = thm "In0_not_In1";
-val In1_not_In0 = thm "In1_not_In0";
-val Un_assoc = thm "Un_assoc";
-val Collect_disj_eq = thm "Collect_disj_eq";
-val empty_def = thm "empty_def";
-val empty_iff = thm "empty_iff";
-
-open DatatypeAux;
-open NominalAtoms;
-
-(** FIXME: DatatypePackage should export this function **)
-
-local
-
-fun dt_recs (DtTFree _) = []
- | dt_recs (DtType (_, dts)) = List.concat (map dt_recs dts)
- | dt_recs (DtRec i) = [i];
-
-fun dt_cases (descr: descr) (_, args, constrs) =
- let
- fun the_bname i = Long_Name.base_name (#1 (valOf (AList.lookup (op =) descr i)));
- val bnames = map the_bname (distinct op = (List.concat (map dt_recs args)));
- in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end;
-
-
-fun induct_cases descr =
- DatatypeProp.indexify_names (List.concat (map (dt_cases descr) (map #2 descr)));
-
-fun exhaust_cases descr i = dt_cases descr (valOf (AList.lookup (op =) descr i));
-
-in
-
-fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
-
-fun mk_case_names_exhausts descr new =
- map (RuleCases.case_names o exhaust_cases descr o #1)
- (List.filter (fn ((_, (name, _, _))) => name mem_string new) descr);
-
-end;
-
-(* theory data *)
-
-type descr = (int * (string * dtyp list * (string * (dtyp list * dtyp) list) list)) list;
-
-type nominal_datatype_info =
- {index : int,
- descr : descr,
- sorts : (string * sort) list,
- rec_names : string list,
- rec_rewrites : thm list,
- induction : thm,
- distinct : thm list,
- inject : thm list};
-
-structure NominalDatatypesData = TheoryDataFun
-(
- type T = nominal_datatype_info Symtab.table;
- val empty = Symtab.empty;
- val copy = I;
- val extend = I;
- fun merge _ tabs : T = Symtab.merge (K true) tabs;
-);
-
-val get_nominal_datatypes = NominalDatatypesData.get;
-val put_nominal_datatypes = NominalDatatypesData.put;
-val map_nominal_datatypes = NominalDatatypesData.map;
-val get_nominal_datatype = Symtab.lookup o get_nominal_datatypes;
-
-
-(**** make datatype info ****)
-
-fun make_dt_info descr sorts induct reccomb_names rec_thms
- (((i, (_, (tname, _, _))), distinct), inject) =
- (tname,
- {index = i,
- descr = descr,
- sorts = sorts,
- rec_names = reccomb_names,
- rec_rewrites = rec_thms,
- induction = induct,
- distinct = distinct,
- inject = inject});
-
-(*******************************)
-
-val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
-
-
-(** simplification procedure for sorting permutations **)
-
-val dj_cp = thm "dj_cp";
-
-fun dest_permT (Type ("fun", [Type ("List.list", [Type ("*", [T, _])]),
- Type ("fun", [_, U])])) = (T, U);
-
-fun permTs_of (Const ("Nominal.perm", T) $ t $ u) = fst (dest_permT T) :: permTs_of u
- | permTs_of _ = [];
-
-fun perm_simproc' thy ss (Const ("Nominal.perm", T) $ t $ (u as Const ("Nominal.perm", U) $ r $ s)) =
- let
- val (aT as Type (a, []), S) = dest_permT T;
- val (bT as Type (b, []), _) = dest_permT U
- in if aT mem permTs_of u andalso aT <> bT then
- let
- val cp = cp_inst_of thy a b;
- val dj = dj_thm_of thy b a;
- val dj_cp' = [cp, dj] MRS dj_cp;
- val cert = SOME o cterm_of thy
- in
- SOME (mk_meta_eq (Drule.instantiate' [SOME (ctyp_of thy S)]
- [cert t, cert r, cert s] dj_cp'))
- end
- else NONE
- end
- | perm_simproc' thy ss _ = NONE;
-
-val perm_simproc =
- Simplifier.simproc (the_context ()) "perm_simp" ["pi1 \<bullet> (pi2 \<bullet> x)"] perm_simproc';
-
-val meta_spec = thm "meta_spec";
-
-fun projections rule =
- ProjectRule.projections (ProofContext.init (Thm.theory_of_thm rule)) rule
- |> map (standard #> RuleCases.save rule);
-
-val supp_prod = thm "supp_prod";
-val fresh_prod = thm "fresh_prod";
-val supports_fresh = thm "supports_fresh";
-val supports_def = thm "Nominal.supports_def";
-val fresh_def = thm "fresh_def";
-val supp_def = thm "supp_def";
-val rev_simps = thms "rev.simps";
-val app_simps = thms "append.simps";
-val at_fin_set_supp = thm "at_fin_set_supp";
-val at_fin_set_fresh = thm "at_fin_set_fresh";
-val abs_fun_eq1 = thm "abs_fun_eq1";
-
-val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq];
-
-fun mk_perm Ts t u =
- let
- val T = fastype_of1 (Ts, t);
- val U = fastype_of1 (Ts, u)
- in Const ("Nominal.perm", T --> U --> U) $ t $ u end;
-
-fun perm_of_pair (x, y) =
- let
- val T = fastype_of x;
- val pT = mk_permT T
- in Const ("List.list.Cons", HOLogic.mk_prodT (T, T) --> pT --> pT) $
- HOLogic.mk_prod (x, y) $ Const ("List.list.Nil", pT)
- end;
-
-fun mk_not_sym ths = maps (fn th => case prop_of th of
- _ $ (Const ("Not", _) $ (Const ("op =", _) $ _ $ _)) => [th, th RS not_sym]
- | _ => [th]) ths;
-
-fun fresh_const T U = Const ("Nominal.fresh", T --> U --> HOLogic.boolT);
-fun fresh_star_const T U =
- Const ("Nominal.fresh_star", HOLogic.mk_setT T --> U --> HOLogic.boolT);
-
-fun gen_add_nominal_datatype prep_typ config new_type_names dts thy =
- let
- (* this theory is used just for parsing *)
-
- val tmp_thy = thy |>
- Theory.copy |>
- Sign.add_types (map (fn (tvs, tname, mx, _) =>
- (Binding.name tname, length tvs, mx)) dts);
-
- val atoms = atoms_of thy;
-
- fun prep_constr ((constrs, sorts), (cname, cargs, mx)) =
- let val (cargs', sorts') = Library.foldl (prep_typ tmp_thy) (([], sorts), cargs)
- in (constrs @ [(cname, cargs', mx)], sorts') end
-
- fun prep_dt_spec ((dts, sorts), (tvs, tname, mx, constrs)) =
- let val (constrs', sorts') = Library.foldl prep_constr (([], sorts), constrs)
- in (dts @ [(tvs, tname, mx, constrs')], sorts') end
-
- val (dts', sorts) = Library.foldl prep_dt_spec (([], []), dts);
- val tyvars = map (map (fn s =>
- (s, the (AList.lookup (op =) sorts s))) o #1) dts';
-
- fun inter_sort thy S S' = Type.inter_sort (Sign.tsig_of thy) (S, S');
- fun augment_sort_typ thy S =
- let val S = Sign.certify_sort thy S
- in map_type_tfree (fn (s, S') => TFree (s,
- if member (op = o apsnd fst) sorts s then inter_sort thy S S' else S'))
- end;
- fun augment_sort thy S = map_types (augment_sort_typ thy S);
-
- val types_syntax = map (fn (tvs, tname, mx, constrs) => (tname, mx)) dts';
- val constr_syntax = map (fn (tvs, tname, mx, constrs) =>
- map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts';
-
- val ps = map (fn (_, n, _, _) =>
- (Sign.full_bname tmp_thy n, Sign.full_bname tmp_thy (n ^ "_Rep"))) dts;
- val rps = map Library.swap ps;
-
- fun replace_types (Type ("Nominal.ABS", [T, U])) =
- Type ("fun", [T, Type ("Nominal.noption", [replace_types U])])
- | replace_types (Type (s, Ts)) =
- Type (getOpt (AList.lookup op = ps s, s), map replace_types Ts)
- | replace_types T = T;
-
- val dts'' = map (fn (tvs, tname, mx, constrs) => (tvs, Binding.name (tname ^ "_Rep"), NoSyn,
- map (fn (cname, cargs, mx) => (Binding.name (cname ^ "_Rep"),
- map replace_types cargs, NoSyn)) constrs)) dts';
-
- val new_type_names' = map (fn n => n ^ "_Rep") new_type_names;
- val full_new_type_names' = map (Sign.full_bname thy) new_type_names';
-
- val ({induction, ...},thy1) =
- DatatypePackage.add_datatype config new_type_names' dts'' thy;
-
- val SOME {descr, ...} = Symtab.lookup
- (DatatypePackage.get_datatypes thy1) (hd full_new_type_names');
- fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
-
- val big_name = space_implode "_" new_type_names;
-
-
- (**** define permutation functions ****)
-
- val permT = mk_permT (TFree ("'x", HOLogic.typeS));
- val pi = Free ("pi", permT);
- val perm_types = map (fn (i, _) =>
- let val T = nth_dtyp i
- in permT --> T --> T end) descr;
- val perm_names' = DatatypeProp.indexify_names (map (fn (i, _) =>
- "perm_" ^ name_of_typ (nth_dtyp i)) descr);
- val perm_names = replicate (length new_type_names) "Nominal.perm" @
- map (Sign.full_bname thy1) (List.drop (perm_names', length new_type_names));
- val perm_names_types = perm_names ~~ perm_types;
- val perm_names_types' = perm_names' ~~ perm_types;
-
- val perm_eqs = maps (fn (i, (_, _, constrs)) =>
- let val T = nth_dtyp i
- in map (fn (cname, dts) =>
- let
- val Ts = map (typ_of_dtyp descr sorts) dts;
- val names = Name.variant_list ["pi"] (DatatypeProp.make_tnames Ts);
- val args = map Free (names ~~ Ts);
- val c = Const (cname, Ts ---> T);
- fun perm_arg (dt, x) =
- let val T = type_of x
- in if is_rec_type dt then
- let val (Us, _) = strip_type T
- in list_abs (map (pair "x") Us,
- Free (nth perm_names_types' (body_index dt)) $ pi $
- list_comb (x, map (fn (i, U) =>
- Const ("Nominal.perm", permT --> U --> U) $
- (Const ("List.rev", permT --> permT) $ pi) $
- Bound i) ((length Us - 1 downto 0) ~~ Us)))
- end
- else Const ("Nominal.perm", permT --> T --> T) $ pi $ x
- end;
- in
- (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq
- (Free (nth perm_names_types' i) $
- Free ("pi", mk_permT (TFree ("'x", HOLogic.typeS))) $
- list_comb (c, args),
- list_comb (c, map perm_arg (dts ~~ args)))))
- end) constrs
- end) descr;
-
- val (perm_simps, thy2) =
- PrimrecPackage.add_primrec_overloaded
- (map (fn (s, sT) => (s, sT, false))
- (List.take (perm_names' ~~ perm_names_types, length new_type_names)))
- (map (fn s => (Binding.name s, NONE, NoSyn)) perm_names') perm_eqs thy1;
-
- (**** prove that permutation functions introduced by unfolding are ****)
- (**** equivalent to already existing permutation functions ****)
-
- val _ = warning ("length descr: " ^ string_of_int (length descr));
- val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names));
-
- val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types);
- val perm_fun_def = PureThy.get_thm thy2 "perm_fun_def";
-
- val unfolded_perm_eq_thms =
- if length descr = length new_type_names then []
- else map standard (List.drop (split_conj_thm
- (Goal.prove_global thy2 [] []
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
- (map (fn (c as (s, T), x) =>
- let val [T1, T2] = binder_types T
- in HOLogic.mk_eq (Const c $ pi $ Free (x, T2),
- Const ("Nominal.perm", T) $ pi $ Free (x, T2))
- end)
- (perm_names_types ~~ perm_indnames))))
- (fn _ => EVERY [indtac induction perm_indnames 1,
- ALLGOALS (asm_full_simp_tac
- (simpset_of thy2 addsimps [perm_fun_def]))])),
- length new_type_names));
-
- (**** prove [] \<bullet> t = t ****)
-
- val _ = warning "perm_empty_thms";
-
- val perm_empty_thms = List.concat (map (fn a =>
- let val permT = mk_permT (Type (a, []))
- in map standard (List.take (split_conj_thm
- (Goal.prove_global thy2 [] []
- (augment_sort thy2 [pt_class_of thy2 a]
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
- (map (fn ((s, T), x) => HOLogic.mk_eq
- (Const (s, permT --> T --> T) $
- Const ("List.list.Nil", permT) $ Free (x, T),
- Free (x, T)))
- (perm_names ~~
- map body_type perm_types ~~ perm_indnames)))))
- (fn _ => EVERY [indtac induction perm_indnames 1,
- ALLGOALS (asm_full_simp_tac (simpset_of thy2))])),
- length new_type_names))
- end)
- atoms);
-
- (**** prove (pi1 @ pi2) \<bullet> t = pi1 \<bullet> (pi2 \<bullet> t) ****)
-
- val _ = warning "perm_append_thms";
-
- (*FIXME: these should be looked up statically*)
- val at_pt_inst = PureThy.get_thm thy2 "at_pt_inst";
- val pt2 = PureThy.get_thm thy2 "pt2";
-
- val perm_append_thms = List.concat (map (fn a =>
- let
- val permT = mk_permT (Type (a, []));
- val pi1 = Free ("pi1", permT);
- val pi2 = Free ("pi2", permT);
- val pt_inst = pt_inst_of thy2 a;
- val pt2' = pt_inst RS pt2;
- val pt2_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "2") a);
- in List.take (map standard (split_conj_thm
- (Goal.prove_global thy2 [] []
- (augment_sort thy2 [pt_class_of thy2 a]
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
- (map (fn ((s, T), x) =>
- let val perm = Const (s, permT --> T --> T)
- in HOLogic.mk_eq
- (perm $ (Const ("List.append", permT --> permT --> permT) $
- pi1 $ pi2) $ Free (x, T),
- perm $ pi1 $ (perm $ pi2 $ Free (x, T)))
- end)
- (perm_names ~~
- map body_type perm_types ~~ perm_indnames)))))
- (fn _ => EVERY [indtac induction perm_indnames 1,
- ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt2', pt2_ax]))]))),
- length new_type_names)
- end) atoms);
-
- (**** prove pi1 ~ pi2 ==> pi1 \<bullet> t = pi2 \<bullet> t ****)
-
- val _ = warning "perm_eq_thms";
-
- val pt3 = PureThy.get_thm thy2 "pt3";
- val pt3_rev = PureThy.get_thm thy2 "pt3_rev";
-
- val perm_eq_thms = List.concat (map (fn a =>
- let
- val permT = mk_permT (Type (a, []));
- val pi1 = Free ("pi1", permT);
- val pi2 = Free ("pi2", permT);
- val at_inst = at_inst_of thy2 a;
- val pt_inst = pt_inst_of thy2 a;
- val pt3' = pt_inst RS pt3;
- val pt3_rev' = at_inst RS (pt_inst RS pt3_rev);
- val pt3_ax = PureThy.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "3") a);
- in List.take (map standard (split_conj_thm
- (Goal.prove_global thy2 [] []
- (augment_sort thy2 [pt_class_of thy2 a] (Logic.mk_implies
- (HOLogic.mk_Trueprop (Const ("Nominal.prm_eq",
- permT --> permT --> HOLogic.boolT) $ pi1 $ pi2),
- HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
- (map (fn ((s, T), x) =>
- let val perm = Const (s, permT --> T --> T)
- in HOLogic.mk_eq
- (perm $ pi1 $ Free (x, T),
- perm $ pi2 $ Free (x, T))
- end)
- (perm_names ~~
- map body_type perm_types ~~ perm_indnames))))))
- (fn _ => EVERY [indtac induction perm_indnames 1,
- ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))),
- length new_type_names)
- end) atoms);
-
- (**** prove pi1 \<bullet> (pi2 \<bullet> t) = (pi1 \<bullet> pi2) \<bullet> (pi1 \<bullet> t) ****)
-
- val cp1 = PureThy.get_thm thy2 "cp1";
- val dj_cp = PureThy.get_thm thy2 "dj_cp";
- val pt_perm_compose = PureThy.get_thm thy2 "pt_perm_compose";
- val pt_perm_compose_rev = PureThy.get_thm thy2 "pt_perm_compose_rev";
- val dj_perm_perm_forget = PureThy.get_thm thy2 "dj_perm_perm_forget";
-
- fun composition_instance name1 name2 thy =
- let
- val cp_class = cp_class_of thy name1 name2;
- val pt_class =
- if name1 = name2 then [pt_class_of thy name1]
- else [];
- val permT1 = mk_permT (Type (name1, []));
- val permT2 = mk_permT (Type (name2, []));
- val Ts = map body_type perm_types;
- val cp_inst = cp_inst_of thy name1 name2;
- val simps = simpset_of thy addsimps (perm_fun_def ::
- (if name1 <> name2 then
- let val dj = dj_thm_of thy name2 name1
- in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end
- else
- let
- val at_inst = at_inst_of thy name1;
- val pt_inst = pt_inst_of thy name1;
- in
- [cp_inst RS cp1 RS sym,
- at_inst RS (pt_inst RS pt_perm_compose) RS sym,
- at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym]
- end))
- val sort = Sign.certify_sort thy (cp_class :: pt_class);
- val thms = split_conj_thm (Goal.prove_global thy [] []
- (augment_sort thy sort
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
- (map (fn ((s, T), x) =>
- let
- val pi1 = Free ("pi1", permT1);
- val pi2 = Free ("pi2", permT2);
- val perm1 = Const (s, permT1 --> T --> T);
- val perm2 = Const (s, permT2 --> T --> T);
- val perm3 = Const ("Nominal.perm", permT1 --> permT2 --> permT2)
- in HOLogic.mk_eq
- (perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)),
- perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T)))
- end)
- (perm_names ~~ Ts ~~ perm_indnames)))))
- (fn _ => EVERY [indtac induction perm_indnames 1,
- ALLGOALS (asm_full_simp_tac simps)]))
- in
- fold (fn (s, tvs) => fn thy => AxClass.prove_arity
- (s, map (inter_sort thy sort o snd) tvs, [cp_class])
- (Class.intro_classes_tac [] THEN ALLGOALS (resolve_tac thms)) thy)
- (full_new_type_names' ~~ tyvars) thy
- end;
-
- val (perm_thmss,thy3) = thy2 |>
- fold (fn name1 => fold (composition_instance name1) atoms) atoms |>
- fold (fn atom => fn thy =>
- let val pt_name = pt_class_of thy atom
- in
- fold (fn (s, tvs) => fn thy => AxClass.prove_arity
- (s, map (inter_sort thy [pt_name] o snd) tvs, [pt_name])
- (EVERY
- [Class.intro_classes_tac [],
- resolve_tac perm_empty_thms 1,
- resolve_tac perm_append_thms 1,
- resolve_tac perm_eq_thms 1, assume_tac 1]) thy)
- (full_new_type_names' ~~ tyvars) thy
- end) atoms |>
- PureThy.add_thmss
- [((Binding.name (space_implode "_" new_type_names ^ "_unfolded_perm_eq"),
- unfolded_perm_eq_thms), [Simplifier.simp_add]),
- ((Binding.name (space_implode "_" new_type_names ^ "_perm_empty"),
- perm_empty_thms), [Simplifier.simp_add]),
- ((Binding.name (space_implode "_" new_type_names ^ "_perm_append"),
- perm_append_thms), [Simplifier.simp_add]),
- ((Binding.name (space_implode "_" new_type_names ^ "_perm_eq"),
- perm_eq_thms), [Simplifier.simp_add])];
-
- (**** Define representing sets ****)
-
- val _ = warning "representing sets";
-
- val rep_set_names = DatatypeProp.indexify_names
- (map (fn (i, _) => name_of_typ (nth_dtyp i) ^ "_set") descr);
- val big_rep_name =
- space_implode "_" (DatatypeProp.indexify_names (List.mapPartial
- (fn (i, ("Nominal.noption", _, _)) => NONE
- | (i, _) => SOME (name_of_typ (nth_dtyp i))) descr)) ^ "_set";
- val _ = warning ("big_rep_name: " ^ big_rep_name);
-
- fun strip_option (dtf as DtType ("fun", [dt, DtRec i])) =
- (case AList.lookup op = descr i of
- SOME ("Nominal.noption", _, [(_, [dt']), _]) =>
- apfst (cons dt) (strip_option dt')
- | _ => ([], dtf))
- | strip_option (DtType ("fun", [dt, DtType ("Nominal.noption", [dt'])])) =
- apfst (cons dt) (strip_option dt')
- | strip_option dt = ([], dt);
-
- val dt_atomTs = distinct op = (map (typ_of_dtyp descr sorts)
- (List.concat (map (fn (_, (_, _, cs)) => List.concat
- (map (List.concat o map (fst o strip_option) o snd) cs)) descr)));
- val dt_atoms = map (fst o dest_Type) dt_atomTs;
-
- fun make_intr s T (cname, cargs) =
- let
- fun mk_prem (dt, (j, j', prems, ts)) =
- let
- val (dts, dt') = strip_option dt;
- val (dts', dt'') = strip_dtyp dt';
- val Ts = map (typ_of_dtyp descr sorts) dts;
- val Us = map (typ_of_dtyp descr sorts) dts';
- val T = typ_of_dtyp descr sorts dt'';
- val free = mk_Free "x" (Us ---> T) j;
- val free' = app_bnds free (length Us);
- fun mk_abs_fun (T, (i, t)) =
- let val U = fastype_of t
- in (i + 1, Const ("Nominal.abs_fun", [T, U, T] --->
- Type ("Nominal.noption", [U])) $ mk_Free "y" T i $ t)
- end
- in (j + 1, j' + length Ts,
- case dt'' of
- DtRec k => list_all (map (pair "x") Us,
- HOLogic.mk_Trueprop (Free (List.nth (rep_set_names, k),
- T --> HOLogic.boolT) $ free')) :: prems
- | _ => prems,
- snd (List.foldr mk_abs_fun (j', free) Ts) :: ts)
- end;
-
- val (_, _, prems, ts) = List.foldr mk_prem (1, 1, [], []) cargs;
- val concl = HOLogic.mk_Trueprop (Free (s, T --> HOLogic.boolT) $
- list_comb (Const (cname, map fastype_of ts ---> T), ts))
- in Logic.list_implies (prems, concl)
- end;
-
- val (intr_ts, (rep_set_names', recTs')) =
- apfst List.concat (apsnd ListPair.unzip (ListPair.unzip (List.mapPartial
- (fn ((_, ("Nominal.noption", _, _)), _) => NONE
- | ((i, (_, _, constrs)), rep_set_name) =>
- let val T = nth_dtyp i
- in SOME (map (make_intr rep_set_name T) constrs,
- (rep_set_name, T))
- end)
- (descr ~~ rep_set_names))));
- val rep_set_names'' = map (Sign.full_bname thy3) rep_set_names';
-
- val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy4) =
- InductivePackage.add_inductive_global (serial_string ())
- {quiet_mode = false, verbose = false, kind = Thm.internalK,
- alt_name = Binding.name big_rep_name, coind = false, no_elim = true, no_ind = false,
- skip_mono = true, fork_mono = false}
- (map (fn (s, T) => ((Binding.name s, T --> HOLogic.boolT), NoSyn))
- (rep_set_names' ~~ recTs'))
- [] (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] thy3;
-
- (**** Prove that representing set is closed under permutation ****)
-
- val _ = warning "proving closure under permutation...";
-
- val abs_perm = PureThy.get_thms thy4 "abs_perm";
-
- val perm_indnames' = List.mapPartial
- (fn (x, (_, ("Nominal.noption", _, _))) => NONE | (x, _) => SOME x)
- (perm_indnames ~~ descr);
-
- fun mk_perm_closed name = map (fn th => standard (th RS mp))
- (List.take (split_conj_thm (Goal.prove_global thy4 [] []
- (augment_sort thy4
- (pt_class_of thy4 name :: map (cp_class_of thy4 name) (dt_atoms \ name))
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map
- (fn ((s, T), x) =>
- let
- val S = Const (s, T --> HOLogic.boolT);
- val permT = mk_permT (Type (name, []))
- in HOLogic.mk_imp (S $ Free (x, T),
- S $ (Const ("Nominal.perm", permT --> T --> T) $
- Free ("pi", permT) $ Free (x, T)))
- end) (rep_set_names'' ~~ recTs' ~~ perm_indnames')))))
- (fn _ => EVERY
- [indtac rep_induct [] 1,
- ALLGOALS (simp_tac (simpset_of thy4 addsimps
- (symmetric perm_fun_def :: abs_perm))),
- ALLGOALS (resolve_tac rep_intrs THEN_ALL_NEW assume_tac)])),
- length new_type_names));
-
- val perm_closed_thmss = map mk_perm_closed atoms;
-
- (**** typedef ****)
-
- val _ = warning "defining type...";
-
- val (typedefs, thy6) =
- thy4
- |> fold_map (fn ((((name, mx), tvs), (cname, U)), name') => fn thy =>
- TypedefPackage.add_typedef false (SOME (Binding.name name'))
- (Binding.name name, map fst tvs, mx)
- (Const ("Collect", (U --> HOLogic.boolT) --> HOLogic.mk_setT U) $
- Const (cname, U --> HOLogic.boolT)) NONE
- (rtac exI 1 THEN rtac CollectI 1 THEN
- QUIET_BREADTH_FIRST (has_fewer_prems 1)
- (resolve_tac rep_intrs 1)) thy |> (fn ((_, r), thy) =>
- let
- val permT = mk_permT
- (TFree (Name.variant (map fst tvs) "'a", HOLogic.typeS));
- val pi = Free ("pi", permT);
- val T = Type (Sign.intern_type thy name, map TFree tvs);
- in apfst (pair r o hd)
- (PureThy.add_defs_unchecked true [((Binding.name ("prm_" ^ name ^ "_def"), Logic.mk_equals
- (Const ("Nominal.perm", permT --> T --> T) $ pi $ Free ("x", T),
- Const (Sign.intern_const thy ("Abs_" ^ name), U --> T) $
- (Const ("Nominal.perm", permT --> U --> U) $ pi $
- (Const (Sign.intern_const thy ("Rep_" ^ name), T --> U) $
- Free ("x", T))))), [])] thy)
- end))
- (types_syntax ~~ tyvars ~~
- List.take (rep_set_names'' ~~ recTs', length new_type_names) ~~
- new_type_names);
-
- val perm_defs = map snd typedefs;
- val Abs_inverse_thms = map (collect_simp o #Abs_inverse o fst) typedefs;
- val Rep_inverse_thms = map (#Rep_inverse o fst) typedefs;
- val Rep_thms = map (collect_simp o #Rep o fst) typedefs;
-
-
- (** prove that new types are in class pt_<name> **)
-
- val _ = warning "prove that new types are in class pt_<name> ...";
-
- fun pt_instance (atom, perm_closed_thms) =
- fold (fn ((((((Abs_inverse, Rep_inverse), Rep),
- perm_def), name), tvs), perm_closed) => fn thy =>
- let
- val pt_class = pt_class_of thy atom;
- val sort = Sign.certify_sort thy
- (pt_class :: map (cp_class_of thy atom) (dt_atoms \ atom))
- in AxClass.prove_arity
- (Sign.intern_type thy name,
- map (inter_sort thy sort o snd) tvs, [pt_class])
- (EVERY [Class.intro_classes_tac [],
- rewrite_goals_tac [perm_def],
- asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1,
- asm_full_simp_tac (simpset_of thy addsimps
- [Rep RS perm_closed RS Abs_inverse]) 1,
- asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy
- ("pt_" ^ Long_Name.base_name atom ^ "3")]) 1]) thy
- end)
- (Abs_inverse_thms ~~ Rep_inverse_thms ~~ Rep_thms ~~ perm_defs ~~
- new_type_names ~~ tyvars ~~ perm_closed_thms);
-
-
- (** prove that new types are in class cp_<name1>_<name2> **)
-
- val _ = warning "prove that new types are in class cp_<name1>_<name2> ...";
-
- fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy =
- let
- val cp_class = cp_class_of thy atom1 atom2;
- val sort = Sign.certify_sort thy
- (pt_class_of thy atom1 :: map (cp_class_of thy atom1) (dt_atoms \ atom1) @
- (if atom1 = atom2 then [cp_class_of thy atom1 atom1] else
- pt_class_of thy atom2 :: map (cp_class_of thy atom2) (dt_atoms \ atom2)));
- val cp1' = cp_inst_of thy atom1 atom2 RS cp1
- in fold (fn ((((((Abs_inverse, Rep),
- perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy =>
- AxClass.prove_arity
- (Sign.intern_type thy name,
- map (inter_sort thy sort o snd) tvs, [cp_class])
- (EVERY [Class.intro_classes_tac [],
- rewrite_goals_tac [perm_def],
- asm_full_simp_tac (simpset_of thy addsimps
- ((Rep RS perm_closed1 RS Abs_inverse) ::
- (if atom1 = atom2 then []
- else [Rep RS perm_closed2 RS Abs_inverse]))) 1,
- cong_tac 1,
- rtac refl 1,
- rtac cp1' 1]) thy)
- (Abs_inverse_thms ~~ Rep_thms ~~ perm_defs ~~ new_type_names ~~
- tyvars ~~ perm_closed_thms1 ~~ perm_closed_thms2) thy
- end;
-
- val thy7 = fold (fn x => fn thy => thy |>
- pt_instance x |>
- fold (cp_instance x) (atoms ~~ perm_closed_thmss))
- (atoms ~~ perm_closed_thmss) thy6;
-
- (**** constructors ****)
-
- fun mk_abs_fun (x, t) =
- let
- val T = fastype_of x;
- val U = fastype_of t
- in
- Const ("Nominal.abs_fun", T --> U --> T -->
- Type ("Nominal.noption", [U])) $ x $ t
- end;
-
- val (ty_idxs, _) = List.foldl
- (fn ((i, ("Nominal.noption", _, _)), p) => p
- | ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr;
-
- fun reindex (DtType (s, dts)) = DtType (s, map reindex dts)
- | reindex (DtRec i) = DtRec (the (AList.lookup op = ty_idxs i))
- | reindex dt = dt;
-
- fun strip_suffix i s = implode (List.take (explode s, size s - i));
-
- (** strips the "_Rep" in type names *)
- fun strip_nth_name i s =
- let val xs = Long_Name.explode s;
- in Long_Name.implode (Library.nth_map (length xs - i) (strip_suffix 4) xs) end;
-
- val (descr'', ndescr) = ListPair.unzip (map_filter
- (fn (i, ("Nominal.noption", _, _)) => NONE
- | (i, (s, dts, constrs)) =>
- let
- val SOME index = AList.lookup op = ty_idxs i;
- val (constrs2, constrs1) =
- map_split (fn (cname, cargs) =>
- apsnd (pair (strip_nth_name 2 (strip_nth_name 1 cname)))
- (fold_map (fn dt => fn dts =>
- let val (dts', dt') = strip_option dt
- in ((length dts, length dts'), dts @ dts' @ [reindex dt']) end)
- cargs [])) constrs
- in SOME ((index, (strip_nth_name 1 s, map reindex dts, constrs1)),
- (index, constrs2))
- end) descr);
-
- val (descr1, descr2) = chop (length new_type_names) descr'';
- val descr' = [descr1, descr2];
-
- fun partition_cargs idxs xs = map (fn (i, j) =>
- (List.take (List.drop (xs, i), j), List.nth (xs, i + j))) idxs;
-
- val pdescr = map (fn ((i, (s, dts, constrs)), (_, idxss)) => (i, (s, dts,
- map (fn ((cname, cargs), idxs) => (cname, partition_cargs idxs cargs))
- (constrs ~~ idxss)))) (descr'' ~~ ndescr);
-
- fun nth_dtyp' i = typ_of_dtyp descr'' sorts (DtRec i);
-
- val rep_names = map (fn s =>
- Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names;
- val abs_names = map (fn s =>
- Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names;
-
- val recTs = get_rec_types descr'' sorts;
- val newTs' = Library.take (length new_type_names, recTs');
- val newTs = Library.take (length new_type_names, recTs);
-
- val full_new_type_names = map (Sign.full_bname thy) new_type_names;
-
- fun make_constr_def tname T T' ((thy, defs, eqns),
- (((cname_rep, _), (cname, cargs)), (cname', mx))) =
- let
- fun constr_arg ((dts, dt), (j, l_args, r_args)) =
- let
- val xs = map (fn (dt, i) => mk_Free "x" (typ_of_dtyp descr'' sorts dt) i)
- (dts ~~ (j upto j + length dts - 1))
- val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
- in
- (j + length dts + 1,
- xs @ x :: l_args,
- List.foldr mk_abs_fun
- (case dt of
- DtRec k => if k < length new_type_names then
- Const (List.nth (rep_names, k), typ_of_dtyp descr'' sorts dt -->
- typ_of_dtyp descr sorts dt) $ x
- else error "nested recursion not (yet) supported"
- | _ => x) xs :: r_args)
- end
-
- val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) cargs;
- val abs_name = Sign.intern_const thy ("Abs_" ^ tname);
- val rep_name = Sign.intern_const thy ("Rep_" ^ tname);
- val constrT = map fastype_of l_args ---> T;
- val lhs = list_comb (Const (cname, constrT), l_args);
- val rhs = list_comb (Const (cname_rep, map fastype_of r_args ---> T'), r_args);
- val def = Logic.mk_equals (lhs, Const (abs_name, T' --> T) $ rhs);
- val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
- (Const (rep_name, T --> T') $ lhs, rhs));
- val def_name = (Long_Name.base_name cname) ^ "_def";
- val ([def_thm], thy') = thy |>
- Sign.add_consts_i [(Binding.name cname', constrT, mx)] |>
- (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)]
- in (thy', defs @ [def_thm], eqns @ [eqn]) end;
-
- fun dt_constr_defs ((thy, defs, eqns, dist_lemmas), ((((((_, (_, _, constrs)),
- (_, (_, _, constrs'))), tname), T), T'), constr_syntax)) =
- let
- val rep_const = cterm_of thy
- (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T'));
- val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
- val (thy', defs', eqns') = Library.foldl (make_constr_def tname T T')
- ((Sign.add_path tname thy, defs, []), constrs ~~ constrs' ~~ constr_syntax)
- in
- (parent_path (#flat_names config) thy', defs', eqns @ [eqns'], dist_lemmas @ [dist])
- end;
-
- val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = Library.foldl dt_constr_defs
- ((thy7, [], [], []), List.take (descr, length new_type_names) ~~
- List.take (pdescr, length new_type_names) ~~
- new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax);
-
- val abs_inject_thms = map (collect_simp o #Abs_inject o fst) typedefs
- val rep_inject_thms = map (#Rep_inject o fst) typedefs
-
- (* prove theorem Rep_i (Constr_j ...) = Constr'_j ... *)
-
- fun prove_constr_rep_thm eqn =
- let
- val inj_thms = map (fn r => r RS iffD1) abs_inject_thms;
- val rewrites = constr_defs @ map mk_meta_eq Rep_inverse_thms
- in Goal.prove_global thy8 [] [] eqn (fn _ => EVERY
- [resolve_tac inj_thms 1,
- rewrite_goals_tac rewrites,
- rtac refl 3,
- resolve_tac rep_intrs 2,
- REPEAT (resolve_tac Rep_thms 1)])
- end;
-
- val constr_rep_thmss = map (map prove_constr_rep_thm) constr_rep_eqns;
-
- (* prove theorem pi \<bullet> Rep_i x = Rep_i (pi \<bullet> x) *)
-
- fun prove_perm_rep_perm (atom, perm_closed_thms) = map (fn th =>
- let
- val _ $ (_ $ (Rep $ x)) = Logic.unvarify (prop_of th);
- val Type ("fun", [T, U]) = fastype_of Rep;
- val permT = mk_permT (Type (atom, []));
- val pi = Free ("pi", permT);
- in
- Goal.prove_global thy8 [] []
- (augment_sort thy8
- (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (dt_atoms \ atom))
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (Const ("Nominal.perm", permT --> U --> U) $ pi $ (Rep $ x),
- Rep $ (Const ("Nominal.perm", permT --> T --> T) $ pi $ x)))))
- (fn _ => simp_tac (HOL_basic_ss addsimps (perm_defs @ Abs_inverse_thms @
- perm_closed_thms @ Rep_thms)) 1)
- end) Rep_thms;
-
- val perm_rep_perm_thms = List.concat (map prove_perm_rep_perm
- (atoms ~~ perm_closed_thmss));
-
- (* prove distinctness theorems *)
-
- val distinct_props = DatatypeProp.make_distincts descr' sorts;
- val dist_rewrites = map2 (fn rep_thms => fn dist_lemma =>
- dist_lemma :: rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0])
- constr_rep_thmss dist_lemmas;
-
- fun prove_distinct_thms _ (_, []) = []
- | prove_distinct_thms (p as (rep_thms, dist_lemma)) (k, t :: ts) =
- let
- val dist_thm = Goal.prove_global thy8 [] [] t (fn _ =>
- simp_tac (simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1)
- in dist_thm :: standard (dist_thm RS not_sym) ::
- prove_distinct_thms p (k, ts)
- end;
-
- val distinct_thms = map2 prove_distinct_thms
- (constr_rep_thmss ~~ dist_lemmas) distinct_props;
-
- (** prove equations for permutation functions **)
-
- val perm_simps' = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
- let val T = nth_dtyp' i
- in List.concat (map (fn (atom, perm_closed_thms) =>
- map (fn ((cname, dts), constr_rep_thm) =>
- let
- val cname = Sign.intern_const thy8
- (Long_Name.append tname (Long_Name.base_name cname));
- val permT = mk_permT (Type (atom, []));
- val pi = Free ("pi", permT);
-
- fun perm t =
- let val T = fastype_of t
- in Const ("Nominal.perm", permT --> T --> T) $ pi $ t end;
-
- fun constr_arg ((dts, dt), (j, l_args, r_args)) =
- let
- val Ts = map (typ_of_dtyp descr'' sorts) dts;
- val xs = map (fn (T, i) => mk_Free "x" T i)
- (Ts ~~ (j upto j + length dts - 1))
- val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
- in
- (j + length dts + 1,
- xs @ x :: l_args,
- map perm (xs @ [x]) @ r_args)
- end
-
- val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) dts;
- val c = Const (cname, map fastype_of l_args ---> T)
- in
- Goal.prove_global thy8 [] []
- (augment_sort thy8
- (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (dt_atoms \ atom))
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (perm (list_comb (c, l_args)), list_comb (c, r_args)))))
- (fn _ => EVERY
- [simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1,
- simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @
- constr_defs @ perm_closed_thms)) 1,
- TRY (simp_tac (HOL_basic_ss addsimps
- (symmetric perm_fun_def :: abs_perm)) 1),
- TRY (simp_tac (HOL_basic_ss addsimps
- (perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @
- perm_closed_thms)) 1)])
- end) (constrs ~~ constr_rep_thms)) (atoms ~~ perm_closed_thmss))
- end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
-
- (** prove injectivity of constructors **)
-
- val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms;
- val alpha = PureThy.get_thms thy8 "alpha";
- val abs_fresh = PureThy.get_thms thy8 "abs_fresh";
-
- val pt_cp_sort =
- map (pt_class_of thy8) dt_atoms @
- maps (fn s => map (cp_class_of thy8 s) (dt_atoms \ s)) dt_atoms;
-
- val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
- let val T = nth_dtyp' i
- in List.mapPartial (fn ((cname, dts), constr_rep_thm) =>
- if null dts then NONE else SOME
- let
- val cname = Sign.intern_const thy8
- (Long_Name.append tname (Long_Name.base_name cname));
-
- fun make_inj ((dts, dt), (j, args1, args2, eqs)) =
- let
- val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1);
- val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
- val ys = map (fn (T, i) => mk_Free "y" T i) Ts_idx;
- val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts);
- val y = mk_Free "y" (typ_of_dtyp descr'' sorts dt) (j + length dts)
- in
- (j + length dts + 1,
- xs @ (x :: args1), ys @ (y :: args2),
- HOLogic.mk_eq
- (List.foldr mk_abs_fun x xs, List.foldr mk_abs_fun y ys) :: eqs)
- end;
-
- val (_, args1, args2, eqs) = List.foldr make_inj (1, [], [], []) dts;
- val Ts = map fastype_of args1;
- val c = Const (cname, Ts ---> T)
- in
- Goal.prove_global thy8 [] []
- (augment_sort thy8 pt_cp_sort
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)),
- foldr1 HOLogic.mk_conj eqs))))
- (fn _ => EVERY
- [asm_full_simp_tac (simpset_of thy8 addsimps (constr_rep_thm ::
- rep_inject_thms')) 1,
- TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def ::
- alpha @ abs_perm @ abs_fresh @ rep_inject_thms @
- perm_rep_perm_thms)) 1)])
- end) (constrs ~~ constr_rep_thms)
- end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
-
- (** equations for support and freshness **)
-
- val (supp_thms, fresh_thms) = ListPair.unzip (map ListPair.unzip
- (map (fn ((((i, (_, _, constrs)), tname), inject_thms'), perm_thms') =>
- let val T = nth_dtyp' i
- in List.concat (map (fn (cname, dts) => map (fn atom =>
- let
- val cname = Sign.intern_const thy8
- (Long_Name.append tname (Long_Name.base_name cname));
- val atomT = Type (atom, []);
-
- fun process_constr ((dts, dt), (j, args1, args2)) =
- let
- val Ts_idx = map (typ_of_dtyp descr'' sorts) dts ~~ (j upto j + length dts - 1);
- val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
- val x = mk_Free "x" (typ_of_dtyp descr'' sorts dt) (j + length dts)
- in
- (j + length dts + 1,
- xs @ (x :: args1), List.foldr mk_abs_fun x xs :: args2)
- end;
-
- val (_, args1, args2) = List.foldr process_constr (1, [], []) dts;
- val Ts = map fastype_of args1;
- val c = list_comb (Const (cname, Ts ---> T), args1);
- fun supp t =
- Const ("Nominal.supp", fastype_of t --> HOLogic.mk_setT atomT) $ t;
- fun fresh t = fresh_const atomT (fastype_of t) $ Free ("a", atomT) $ t;
- val supp_thm = Goal.prove_global thy8 [] []
- (augment_sort thy8 pt_cp_sort
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (supp c,
- if null dts then HOLogic.mk_set atomT []
- else foldr1 (HOLogic.mk_binop @{const_name Un}) (map supp args2)))))
- (fn _ =>
- simp_tac (HOL_basic_ss addsimps (supp_def ::
- Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un ::
- symmetric empty_def :: finite_emptyI :: simp_thms @
- abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1)
- in
- (supp_thm,
- Goal.prove_global thy8 [] [] (augment_sort thy8 pt_cp_sort
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (fresh c,
- if null dts then HOLogic.true_const
- else foldr1 HOLogic.mk_conj (map fresh args2)))))
- (fn _ =>
- simp_tac (HOL_ss addsimps [Un_iff, empty_iff, fresh_def, supp_thm]) 1))
- end) atoms) constrs)
- end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps')));
-
- (**** weak induction theorem ****)
-
- fun mk_indrule_lemma ((prems, concls), (((i, _), T), U)) =
- let
- val Rep_t = Const (List.nth (rep_names, i), T --> U) $
- mk_Free "x" T i;
-
- val Abs_t = Const (List.nth (abs_names, i), U --> T)
-
- in (prems @ [HOLogic.imp $
- (Const (List.nth (rep_set_names'', i), U --> HOLogic.boolT) $ Rep_t) $
- (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
- concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i])
- end;
-
- val (indrule_lemma_prems, indrule_lemma_concls) =
- Library.foldl mk_indrule_lemma (([], []), (descr'' ~~ recTs ~~ recTs'));
-
- val indrule_lemma = Goal.prove_global thy8 [] []
- (Logic.mk_implies
- (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems),
- HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY
- [REPEAT (etac conjE 1),
- REPEAT (EVERY
- [TRY (rtac conjI 1), full_simp_tac (HOL_basic_ss addsimps Rep_inverse_thms) 1,
- etac mp 1, resolve_tac Rep_thms 1])]);
-
- val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
- val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
- map (Free o apfst fst o dest_Var) Ps;
- val indrule_lemma' = cterm_instantiate
- (map (cterm_of thy8) Ps ~~ map (cterm_of thy8) frees) indrule_lemma;
-
- val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms;
-
- val dt_induct_prop = DatatypeProp.make_ind descr' sorts;
- val dt_induct = Goal.prove_global thy8 []
- (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
- (fn {prems, ...} => EVERY
- [rtac indrule_lemma' 1,
- (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
- EVERY (map (fn (prem, r) => (EVERY
- [REPEAT (eresolve_tac Abs_inverse_thms' 1),
- simp_tac (HOL_basic_ss addsimps [symmetric r]) 1,
- DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
- (prems ~~ constr_defs))]);
-
- val case_names_induct = mk_case_names_induct descr'';
-
- (**** prove that new datatypes have finite support ****)
-
- val _ = warning "proving finite support for the new datatype";
-
- val indnames = DatatypeProp.make_tnames recTs;
-
- val abs_supp = PureThy.get_thms thy8 "abs_supp";
- val supp_atm = PureThy.get_thms thy8 "supp_atm";
-
- val finite_supp_thms = map (fn atom =>
- let val atomT = Type (atom, [])
- in map standard (List.take
- (split_conj_thm (Goal.prove_global thy8 [] []
- (augment_sort thy8 (fs_class_of thy8 atom :: pt_cp_sort)
- (HOLogic.mk_Trueprop
- (foldr1 HOLogic.mk_conj (map (fn (s, T) =>
- Const ("Finite_Set.finite", HOLogic.mk_setT atomT --> HOLogic.boolT) $
- (Const ("Nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T)))
- (indnames ~~ recTs)))))
- (fn _ => indtac dt_induct indnames 1 THEN
- ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps
- (abs_supp @ supp_atm @
- PureThy.get_thms thy8 ("fs_" ^ Long_Name.base_name atom ^ "1") @
- List.concat supp_thms))))),
- length new_type_names))
- end) atoms;
-
- val simp_atts = replicate (length new_type_names) [Simplifier.simp_add];
-
- (* Function to add both the simp and eqvt attributes *)
- (* These two attributes are duplicated on all the types in the mutual nominal datatypes *)
-
- val simp_eqvt_atts = replicate (length new_type_names) [Simplifier.simp_add, NominalThmDecls.eqvt_add];
-
- val (_, thy9) = thy8 |>
- Sign.add_path big_name |>
- PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] ||>>
- PureThy.add_thmss [((Binding.name "inducts", projections dt_induct), [case_names_induct])] ||>
- Sign.parent_path ||>>
- DatatypeAux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>>
- DatatypeAux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>>
- DatatypeAux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>>
- DatatypeAux.store_thmss "inject" new_type_names inject_thms ||>>
- DatatypeAux.store_thmss "supp" new_type_names supp_thms ||>>
- DatatypeAux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||>
- fold (fn (atom, ths) => fn thy =>
- let
- val class = fs_class_of thy atom;
- val sort = Sign.certify_sort thy (class :: pt_cp_sort)
- in fold (fn Type (s, Ts) => AxClass.prove_arity
- (s, map (inter_sort thy sort o snd o dest_TFree) Ts, [class])
- (Class.intro_classes_tac [] THEN resolve_tac ths 1)) newTs thy
- end) (atoms ~~ finite_supp_thms);
-
- (**** strong induction theorem ****)
-
- val pnames = if length descr'' = 1 then ["P"]
- else map (fn i => "P" ^ string_of_int i) (1 upto length descr'');
- val ind_sort = if null dt_atomTs then HOLogic.typeS
- else Sign.certify_sort thy9 (map (fs_class_of thy9) dt_atoms);
- val fsT = TFree ("'n", ind_sort);
- val fsT' = TFree ("'n", HOLogic.typeS);
-
- val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T)))
- (DatatypeProp.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs);
-
- fun make_pred fsT i T =
- Free (List.nth (pnames, i), fsT --> T --> HOLogic.boolT);
-
- fun mk_fresh1 xs [] = []
- | mk_fresh1 xs ((y as (_, T)) :: ys) = map (fn x => HOLogic.mk_Trueprop
- (HOLogic.mk_not (HOLogic.mk_eq (Free y, Free x))))
- (filter (fn (_, U) => T = U) (rev xs)) @
- mk_fresh1 (y :: xs) ys;
-
- fun mk_fresh2 xss [] = []
- | mk_fresh2 xss ((p as (ys, _)) :: yss) = List.concat (map (fn y as (_, T) =>
- map (fn (_, x as (_, U)) => HOLogic.mk_Trueprop
- (fresh_const T U $ Free y $ Free x)) (rev xss @ yss)) ys) @
- mk_fresh2 (p :: xss) yss;
-
- fun make_ind_prem fsT f k T ((cname, cargs), idxs) =
- let
- val recs = List.filter is_rec_type cargs;
- val Ts = map (typ_of_dtyp descr'' sorts) cargs;
- val recTs' = map (typ_of_dtyp descr'' sorts) recs;
- val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts);
- val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
- val frees = tnames ~~ Ts;
- val frees' = partition_cargs idxs frees;
- val z = (Name.variant tnames "z", fsT);
-
- fun mk_prem ((dt, s), T) =
- let
- val (Us, U) = strip_type T;
- val l = length Us
- in list_all (z :: map (pair "x") Us, HOLogic.mk_Trueprop
- (make_pred fsT (body_index dt) U $ Bound l $ app_bnds (Free (s, T)) l))
- end;
-
- val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs');
- val prems' = map (fn p as (_, T) => HOLogic.mk_Trueprop
- (f T (Free p) (Free z))) (List.concat (map fst frees')) @
- mk_fresh1 [] (List.concat (map fst frees')) @
- mk_fresh2 [] frees'
-
- in list_all_free (frees @ [z], Logic.list_implies (prems' @ prems,
- HOLogic.mk_Trueprop (make_pred fsT k T $ Free z $
- list_comb (Const (cname, Ts ---> T), map Free frees))))
- end;
-
- val ind_prems = List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) =>
- map (make_ind_prem fsT (fn T => fn t => fn u =>
- fresh_const T fsT $ t $ u) i T)
- (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs));
- val tnames = DatatypeProp.make_tnames recTs;
- val zs = Name.variant_list tnames (replicate (length descr'') "z");
- val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
- (map (fn ((((i, _), T), tname), z) =>
- make_pred fsT i T $ Free (z, fsT) $ Free (tname, T))
- (descr'' ~~ recTs ~~ tnames ~~ zs)));
- val induct = Logic.list_implies (ind_prems, ind_concl);
-
- val ind_prems' =
- map (fn (_, f as Free (_, T)) => list_all_free ([("x", fsT')],
- HOLogic.mk_Trueprop (Const ("Finite_Set.finite",
- (snd (split_last (binder_types T)) --> HOLogic.boolT) -->
- HOLogic.boolT) $ (f $ Free ("x", fsT'))))) fresh_fs @
- List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) =>
- map (make_ind_prem fsT' (fn T => fn t => fn u => HOLogic.Not $
- HOLogic.mk_mem (t, the (AList.lookup op = fresh_fs T) $ u)) i T)
- (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs));
- val ind_concl' = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
- (map (fn ((((i, _), T), tname), z) =>
- make_pred fsT' i T $ Free (z, fsT') $ Free (tname, T))
- (descr'' ~~ recTs ~~ tnames ~~ zs)));
- val induct' = Logic.list_implies (ind_prems', ind_concl');
-
- val aux_ind_vars =
- (DatatypeProp.indexify_names (replicate (length dt_atomTs) "pi") ~~
- map mk_permT dt_atomTs) @ [("z", fsT')];
- val aux_ind_Ts = rev (map snd aux_ind_vars);
- val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
- (map (fn (((i, _), T), tname) =>
- HOLogic.list_all (aux_ind_vars, make_pred fsT' i T $ Bound 0 $
- fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1))
- (Free (tname, T))))
- (descr'' ~~ recTs ~~ tnames)));
-
- val fin_set_supp = map (fn s =>
- at_inst_of thy9 s RS at_fin_set_supp) dt_atoms;
- val fin_set_fresh = map (fn s =>
- at_inst_of thy9 s RS at_fin_set_fresh) dt_atoms;
- val pt1_atoms = map (fn Type (s, _) =>
- PureThy.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "1")) dt_atomTs;
- val pt2_atoms = map (fn Type (s, _) =>
- PureThy.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "2") RS sym) dt_atomTs;
- val exists_fresh' = PureThy.get_thms thy9 "exists_fresh'";
- val fs_atoms = PureThy.get_thms thy9 "fin_supp";
- val abs_supp = PureThy.get_thms thy9 "abs_supp";
- val perm_fresh_fresh = PureThy.get_thms thy9 "perm_fresh_fresh";
- val calc_atm = PureThy.get_thms thy9 "calc_atm";
- val fresh_atm = PureThy.get_thms thy9 "fresh_atm";
- val fresh_left = PureThy.get_thms thy9 "fresh_left";
- val perm_swap = PureThy.get_thms thy9 "perm_swap";
-
- fun obtain_fresh_name' ths ts T (freshs1, freshs2, ctxt) =
- let
- val p = foldr1 HOLogic.mk_prod (ts @ freshs1);
- val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
- (HOLogic.exists_const T $ Abs ("x", T,
- fresh_const T (fastype_of p) $
- Bound 0 $ p)))
- (fn _ => EVERY
- [resolve_tac exists_fresh' 1,
- simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms @
- fin_set_supp @ ths)) 1]);
- val (([cx], ths), ctxt') = Obtain.result
- (fn _ => EVERY
- [etac exE 1,
- full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1,
- REPEAT (etac conjE 1)])
- [ex] ctxt
- in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end;
-
- fun fresh_fresh_inst thy a b =
- let
- val T = fastype_of a;
- val SOME th = find_first (fn th => case prop_of th of
- _ $ (_ $ (Const (_, Type (_, [U, _])) $ _ $ _)) $ _ => U = T
- | _ => false) perm_fresh_fresh
- in
- Drule.instantiate' []
- [SOME (cterm_of thy a), NONE, SOME (cterm_of thy b)] th
- end;
-
- val fs_cp_sort =
- map (fs_class_of thy9) dt_atoms @
- maps (fn s => map (cp_class_of thy9 s) (dt_atoms \ s)) dt_atoms;
-
- (**********************************************************************
- The subgoals occurring in the proof of induct_aux have the
- following parameters:
-
- x_1 ... x_k p_1 ... p_m z
-
- where
-
- x_i : constructor arguments (introduced by weak induction rule)
- p_i : permutations (one for each atom type in the data type)
- z : freshness context
- ***********************************************************************)
-
- val _ = warning "proving strong induction theorem ...";
-
- val induct_aux = Goal.prove_global thy9 []
- (map (augment_sort thy9 fs_cp_sort) ind_prems')
- (augment_sort thy9 fs_cp_sort ind_concl') (fn {prems, context} =>
- let
- val (prems1, prems2) = chop (length dt_atomTs) prems;
- val ind_ss2 = HOL_ss addsimps
- finite_Diff :: abs_fresh @ abs_supp @ fs_atoms;
- val ind_ss1 = ind_ss2 addsimps fresh_left @ calc_atm @
- fresh_atm @ rev_simps @ app_simps;
- val ind_ss3 = HOL_ss addsimps abs_fun_eq1 ::
- abs_perm @ calc_atm @ perm_swap;
- val ind_ss4 = HOL_basic_ss addsimps fresh_left @ prems1 @
- fin_set_fresh @ calc_atm;
- val ind_ss5 = HOL_basic_ss addsimps pt1_atoms;
- val ind_ss6 = HOL_basic_ss addsimps flat perm_simps';
- val th = Goal.prove context [] []
- (augment_sort thy9 fs_cp_sort aux_ind_concl)
- (fn {context = context1, ...} =>
- EVERY (indtac dt_induct tnames 1 ::
- maps (fn ((_, (_, _, constrs)), (_, constrs')) =>
- map (fn ((cname, cargs), is) =>
- REPEAT (rtac allI 1) THEN
- SUBPROOF (fn {prems = iprems, params, concl,
- context = context2, ...} =>
- let
- val concl' = term_of concl;
- val _ $ (_ $ _ $ u) = concl';
- val U = fastype_of u;
- val (xs, params') =
- chop (length cargs) (map term_of params);
- val Ts = map fastype_of xs;
- val cnstr = Const (cname, Ts ---> U);
- val (pis, z) = split_last params';
- val mk_pi = fold_rev (mk_perm []) pis;
- val xs' = partition_cargs is xs;
- val xs'' = map (fn (ts, u) => (map mk_pi ts, mk_pi u)) xs';
- val ts = maps (fn (ts, u) => ts @ [u]) xs'';
- val (freshs1, freshs2, context3) = fold (fn t =>
- let val T = fastype_of t
- in obtain_fresh_name' prems1
- (the (AList.lookup op = fresh_fs T) $ z :: ts) T
- end) (maps fst xs') ([], [], context2);
- val freshs1' = unflat (map fst xs') freshs1;
- val freshs2' = map (Simplifier.simplify ind_ss4)
- (mk_not_sym freshs2);
- val ind_ss1' = ind_ss1 addsimps freshs2';
- val ind_ss3' = ind_ss3 addsimps freshs2';
- val rename_eq =
- if forall (null o fst) xs' then []
- else [Goal.prove context3 [] []
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (list_comb (cnstr, ts),
- list_comb (cnstr, maps (fn ((bs, t), cs) =>
- cs @ [fold_rev (mk_perm []) (map perm_of_pair
- (bs ~~ cs)) t]) (xs'' ~~ freshs1')))))
- (fn _ => EVERY
- (simp_tac (HOL_ss addsimps flat inject_thms) 1 ::
- REPEAT (FIRSTGOAL (rtac conjI)) ::
- maps (fn ((bs, t), cs) =>
- if null bs then []
- else rtac sym 1 :: maps (fn (b, c) =>
- [rtac trans 1, rtac sym 1,
- rtac (fresh_fresh_inst thy9 b c) 1,
- simp_tac ind_ss1' 1,
- simp_tac ind_ss2 1,
- simp_tac ind_ss3' 1]) (bs ~~ cs))
- (xs'' ~~ freshs1')))];
- val th = Goal.prove context3 [] [] concl' (fn _ => EVERY
- [simp_tac (ind_ss6 addsimps rename_eq) 1,
- cut_facts_tac iprems 1,
- (resolve_tac prems THEN_ALL_NEW
- SUBGOAL (fn (t, i) => case Logic.strip_assums_concl t of
- _ $ (Const ("Nominal.fresh", _) $ _ $ _) =>
- simp_tac ind_ss1' i
- | _ $ (Const ("Not", _) $ _) =>
- resolve_tac freshs2' i
- | _ => asm_simp_tac (HOL_basic_ss addsimps
- pt2_atoms addsimprocs [perm_simproc]) i)) 1])
- val final = ProofContext.export context3 context2 [th]
- in
- resolve_tac final 1
- end) context1 1) (constrs ~~ constrs')) (descr'' ~~ ndescr)))
- in
- EVERY
- [cut_facts_tac [th] 1,
- REPEAT (eresolve_tac [conjE, @{thm allE_Nil}] 1),
- REPEAT (etac allE 1),
- REPEAT (TRY (rtac conjI 1) THEN asm_full_simp_tac ind_ss5 1)]
- end);
-
- val induct_aux' = Thm.instantiate ([],
- map (fn (s, v as Var (_, T)) =>
- (cterm_of thy9 v, cterm_of thy9 (Free (s, T))))
- (pnames ~~ map head_of (HOLogic.dest_conj
- (HOLogic.dest_Trueprop (concl_of induct_aux)))) @
- map (fn (_, f) =>
- let val f' = Logic.varify f
- in (cterm_of thy9 f',
- cterm_of thy9 (Const ("Nominal.supp", fastype_of f')))
- end) fresh_fs) induct_aux;
-
- val induct = Goal.prove_global thy9 []
- (map (augment_sort thy9 fs_cp_sort) ind_prems)
- (augment_sort thy9 fs_cp_sort ind_concl)
- (fn {prems, ...} => EVERY
- [rtac induct_aux' 1,
- REPEAT (resolve_tac fs_atoms 1),
- REPEAT ((resolve_tac prems THEN_ALL_NEW
- (etac meta_spec ORELSE' full_simp_tac (HOL_basic_ss addsimps [fresh_def]))) 1)])
-
- val (_, thy10) = thy9 |>
- Sign.add_path big_name |>
- PureThy.add_thms [((Binding.name "strong_induct'", induct_aux), [])] ||>>
- PureThy.add_thms [((Binding.name "strong_induct", induct), [case_names_induct])] ||>>
- PureThy.add_thmss [((Binding.name "strong_inducts", projections induct), [case_names_induct])];
-
- (**** recursion combinator ****)
-
- val _ = warning "defining recursion combinator ...";
-
- val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
-
- val (rec_result_Ts', rec_fn_Ts') = DatatypeProp.make_primrec_Ts descr' sorts used;
-
- val rec_sort = if null dt_atomTs then HOLogic.typeS else
- Sign.certify_sort thy10 pt_cp_sort;
-
- val rec_result_Ts = map (fn TFree (s, _) => TFree (s, rec_sort)) rec_result_Ts';
- val rec_fn_Ts = map (typ_subst_atomic (rec_result_Ts' ~~ rec_result_Ts)) rec_fn_Ts';
-
- val rec_set_Ts = map (fn (T1, T2) =>
- rec_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
-
- val big_rec_name = big_name ^ "_rec_set";
- val rec_set_names' =
- if length descr'' = 1 then [big_rec_name] else
- map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int)
- (1 upto (length descr''));
- val rec_set_names = map (Sign.full_bname thy10) rec_set_names';
-
- val rec_fns = map (uncurry (mk_Free "f"))
- (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts)));
- val rec_sets' = map (fn c => list_comb (Free c, rec_fns))
- (rec_set_names' ~~ rec_set_Ts);
- val rec_sets = map (fn c => list_comb (Const c, rec_fns))
- (rec_set_names ~~ rec_set_Ts);
-
- (* introduction rules for graph of recursion function *)
-
- val rec_preds = map (fn (a, T) =>
- Free (a, T --> HOLogic.boolT)) (pnames ~~ rec_result_Ts);
-
- fun mk_fresh3 rs [] = []
- | mk_fresh3 rs ((p as (ys, z)) :: yss) = List.concat (map (fn y as (_, T) =>
- List.mapPartial (fn ((_, (_, x)), r as (_, U)) => if z = x then NONE
- else SOME (HOLogic.mk_Trueprop
- (fresh_const T U $ Free y $ Free r))) rs) ys) @
- mk_fresh3 rs yss;
-
- (* FIXME: avoid collisions with other variable names? *)
- val rec_ctxt = Free ("z", fsT');
-
- fun make_rec_intr T p rec_set ((rec_intr_ts, rec_prems, rec_prems',
- rec_eq_prems, l), ((cname, cargs), idxs)) =
- let
- val Ts = map (typ_of_dtyp descr'' sorts) cargs;
- val frees = map (fn i => "x" ^ string_of_int i) (1 upto length Ts) ~~ Ts;
- val frees' = partition_cargs idxs frees;
- val binders = List.concat (map fst frees');
- val atomTs = distinct op = (maps (map snd o fst) frees');
- val recs = List.mapPartial
- (fn ((_, DtRec i), p) => SOME (i, p) | _ => NONE)
- (partition_cargs idxs cargs ~~ frees');
- val frees'' = map (fn i => "y" ^ string_of_int i) (1 upto length recs) ~~
- map (fn (i, _) => List.nth (rec_result_Ts, i)) recs;
- val prems1 = map (fn ((i, (_, x)), y) => HOLogic.mk_Trueprop
- (List.nth (rec_sets', i) $ Free x $ Free y)) (recs ~~ frees'');
- val prems2 =
- map (fn f => map (fn p as (_, T) => HOLogic.mk_Trueprop
- (fresh_const T (fastype_of f) $ Free p $ f)) binders) rec_fns;
- val prems3 = mk_fresh1 [] binders @ mk_fresh2 [] frees';
- val prems4 = map (fn ((i, _), y) =>
- HOLogic.mk_Trueprop (List.nth (rec_preds, i) $ Free y)) (recs ~~ frees'');
- val prems5 = mk_fresh3 (recs ~~ frees'') frees';
- val prems6 = maps (fn aT => map (fn y as (_, T) => HOLogic.mk_Trueprop
- (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
- (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ Free y)))
- frees'') atomTs;
- val prems7 = map (fn x as (_, T) => HOLogic.mk_Trueprop
- (fresh_const T fsT' $ Free x $ rec_ctxt)) binders;
- val result = list_comb (List.nth (rec_fns, l), map Free (frees @ frees''));
- val result_freshs = map (fn p as (_, T) =>
- fresh_const T (fastype_of result) $ Free p $ result) binders;
- val P = HOLogic.mk_Trueprop (p $ result)
- in
- (rec_intr_ts @ [Logic.list_implies (List.concat prems2 @ prems3 @ prems1,
- HOLogic.mk_Trueprop (rec_set $
- list_comb (Const (cname, Ts ---> T), map Free frees) $ result))],
- rec_prems @ [list_all_free (frees @ frees'', Logic.list_implies (prems4, P))],
- rec_prems' @ map (fn fr => list_all_free (frees @ frees'',
- Logic.list_implies (List.nth (prems2, l) @ prems3 @ prems5 @ prems7 @ prems6 @ [P],
- HOLogic.mk_Trueprop fr))) result_freshs,
- rec_eq_prems @ [List.concat prems2 @ prems3],
- l + 1)
- end;
-
- val (rec_intr_ts, rec_prems, rec_prems', rec_eq_prems, _) =
- Library.foldl (fn (x, ((((d, d'), T), p), rec_set)) =>
- Library.foldl (make_rec_intr T p rec_set) (x, #3 (snd d) ~~ snd d'))
- (([], [], [], [], 0), descr'' ~~ ndescr ~~ recTs ~~ rec_preds ~~ rec_sets');
-
- val ({intrs = rec_intrs, elims = rec_elims, raw_induct = rec_induct, ...}, thy11) =
- thy10 |>
- InductivePackage.add_inductive_global (serial_string ())
- {quiet_mode = #quiet config, verbose = false, kind = Thm.internalK,
- alt_name = Binding.name big_rec_name, coind = false, no_elim = false, no_ind = false,
- skip_mono = true, fork_mono = false}
- (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
- (map dest_Free rec_fns)
- (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] ||>
- PureThy.hide_fact true (Long_Name.append (Sign.full_bname thy10 big_rec_name) "induct");
-
- (** equivariance **)
-
- val fresh_bij = PureThy.get_thms thy11 "fresh_bij";
- val perm_bij = PureThy.get_thms thy11 "perm_bij";
-
- val (rec_equiv_thms, rec_equiv_thms') = ListPair.unzip (map (fn aT =>
- let
- val permT = mk_permT aT;
- val pi = Free ("pi", permT);
- val rec_fns_pi = map (mk_perm [] pi o uncurry (mk_Free "f"))
- (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts)));
- val rec_sets_pi = map (fn c => list_comb (Const c, rec_fns_pi))
- (rec_set_names ~~ rec_set_Ts);
- val ps = map (fn ((((T, U), R), R'), i) =>
- let
- val x = Free ("x" ^ string_of_int i, T);
- val y = Free ("y" ^ string_of_int i, U)
- in
- (R $ x $ y, R' $ mk_perm [] pi x $ mk_perm [] pi y)
- end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ rec_sets_pi ~~ (1 upto length recTs));
- val ths = map (fn th => standard (th RS mp)) (split_conj_thm
- (Goal.prove_global thy11 [] []
- (augment_sort thy1 pt_cp_sort
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps))))
- (fn _ => rtac rec_induct 1 THEN REPEAT
- (simp_tac (Simplifier.theory_context thy11 HOL_basic_ss
- addsimps flat perm_simps'
- addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
- (resolve_tac rec_intrs THEN_ALL_NEW
- asm_simp_tac (HOL_ss addsimps (fresh_bij @ perm_bij))) 1))))
- val ths' = map (fn ((P, Q), th) =>
- Goal.prove_global thy11 [] []
- (augment_sort thy1 pt_cp_sort
- (Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P)))
- (fn _ => dtac (Thm.instantiate ([],
- [(cterm_of thy11 (Var (("pi", 0), permT)),
- cterm_of thy11 (Const ("List.rev", permT --> permT) $ pi))]) th) 1 THEN
- NominalPermeq.perm_simp_tac HOL_ss 1)) (ps ~~ ths)
- in (ths, ths') end) dt_atomTs);
-
- (** finite support **)
-
- val rec_fin_supp_thms = map (fn aT =>
- let
- val name = Long_Name.base_name (fst (dest_Type aT));
- val fs_name = PureThy.get_thm thy11 ("fs_" ^ name ^ "1");
- val aset = HOLogic.mk_setT aT;
- val finite = Const ("Finite_Set.finite", aset --> HOLogic.boolT);
- val fins = map (fn (f, T) => HOLogic.mk_Trueprop
- (finite $ (Const ("Nominal.supp", T --> aset) $ f)))
- (rec_fns ~~ rec_fn_Ts)
- in
- map (fn th => standard (th RS mp)) (split_conj_thm
- (Goal.prove_global thy11 []
- (map (augment_sort thy11 fs_cp_sort) fins)
- (augment_sort thy11 fs_cp_sort
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
- (map (fn (((T, U), R), i) =>
- let
- val x = Free ("x" ^ string_of_int i, T);
- val y = Free ("y" ^ string_of_int i, U)
- in
- HOLogic.mk_imp (R $ x $ y,
- finite $ (Const ("Nominal.supp", U --> aset) $ y))
- end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~
- (1 upto length recTs))))))
- (fn {prems = fins, ...} =>
- (rtac rec_induct THEN_ALL_NEW cut_facts_tac fins) 1 THEN REPEAT
- (NominalPermeq.finite_guess_tac (HOL_ss addsimps [fs_name]) 1))))
- end) dt_atomTs;
-
- (** freshness **)
-
- val finite_premss = map (fn aT =>
- map (fn (f, T) => HOLogic.mk_Trueprop
- (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
- (Const ("Nominal.supp", T --> HOLogic.mk_setT aT) $ f)))
- (rec_fns ~~ rec_fn_Ts)) dt_atomTs;
-
- val rec_fns' = map (augment_sort thy11 fs_cp_sort) rec_fns;
-
- val rec_fresh_thms = map (fn ((aT, eqvt_ths), finite_prems) =>
- let
- val name = Long_Name.base_name (fst (dest_Type aT));
- val fs_name = PureThy.get_thm thy11 ("fs_" ^ name ^ "1");
- val a = Free ("a", aT);
- val freshs = map (fn (f, fT) => HOLogic.mk_Trueprop
- (fresh_const aT fT $ a $ f)) (rec_fns ~~ rec_fn_Ts)
- in
- map (fn (((T, U), R), eqvt_th) =>
- let
- val x = Free ("x", augment_sort_typ thy11 fs_cp_sort T);
- val y = Free ("y", U);
- val y' = Free ("y'", U)
- in
- standard (Goal.prove (ProofContext.init thy11) []
- (map (augment_sort thy11 fs_cp_sort)
- (finite_prems @
- [HOLogic.mk_Trueprop (R $ x $ y),
- HOLogic.mk_Trueprop (HOLogic.mk_all ("y'", U,
- HOLogic.mk_imp (R $ x $ y', HOLogic.mk_eq (y', y)))),
- HOLogic.mk_Trueprop (fresh_const aT T $ a $ x)] @
- freshs))
- (HOLogic.mk_Trueprop (fresh_const aT U $ a $ y))
- (fn {prems, context} =>
- let
- val (finite_prems, rec_prem :: unique_prem ::
- fresh_prems) = chop (length finite_prems) prems;
- val unique_prem' = unique_prem RS spec RS mp;
- val unique = [unique_prem', unique_prem' RS sym] MRS trans;
- val _ $ (_ $ (_ $ S $ _)) $ _ = prop_of supports_fresh;
- val tuple = foldr1 HOLogic.mk_prod (x :: rec_fns')
- in EVERY
- [rtac (Drule.cterm_instantiate
- [(cterm_of thy11 S,
- cterm_of thy11 (Const ("Nominal.supp",
- fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))]
- supports_fresh) 1,
- simp_tac (HOL_basic_ss addsimps
- [supports_def, symmetric fresh_def, fresh_prod]) 1,
- REPEAT_DETERM (resolve_tac [allI, impI] 1),
- REPEAT_DETERM (etac conjE 1),
- rtac unique 1,
- SUBPROOF (fn {prems = prems', params = [a, b], ...} => EVERY
- [cut_facts_tac [rec_prem] 1,
- rtac (Thm.instantiate ([],
- [(cterm_of thy11 (Var (("pi", 0), mk_permT aT)),
- cterm_of thy11 (perm_of_pair (term_of a, term_of b)))]) eqvt_th) 1,
- asm_simp_tac (HOL_ss addsimps
- (prems' @ perm_swap @ perm_fresh_fresh)) 1]) context 1,
- rtac rec_prem 1,
- simp_tac (HOL_ss addsimps (fs_name ::
- supp_prod :: finite_Un :: finite_prems)) 1,
- simp_tac (HOL_ss addsimps (symmetric fresh_def ::
- fresh_prod :: fresh_prems)) 1]
- end))
- end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ eqvt_ths)
- end) (dt_atomTs ~~ rec_equiv_thms' ~~ finite_premss);
-
- (** uniqueness **)
-
- val fun_tuple = foldr1 HOLogic.mk_prod (rec_ctxt :: rec_fns);
- val fun_tupleT = fastype_of fun_tuple;
- val rec_unique_frees =
- DatatypeProp.indexify_names (replicate (length recTs) "x") ~~ recTs;
- val rec_unique_frees'' = map (fn (s, T) => (s ^ "'", T)) rec_unique_frees;
- val rec_unique_frees' =
- DatatypeProp.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts;
- val rec_unique_concls = map (fn ((x, U), R) =>
- Const ("Ex1", (U --> HOLogic.boolT) --> HOLogic.boolT) $
- Abs ("y", U, R $ Free x $ Bound 0))
- (rec_unique_frees ~~ rec_result_Ts ~~ rec_sets);
-
- val induct_aux_rec = Drule.cterm_instantiate
- (map (pairself (cterm_of thy11) o apsnd (augment_sort thy11 fs_cp_sort))
- (map (fn (aT, f) => (Logic.varify f, Abs ("z", HOLogic.unitT,
- Const ("Nominal.supp", fun_tupleT --> HOLogic.mk_setT aT) $ fun_tuple)))
- fresh_fs @
- map (fn (((P, T), (x, U)), Q) =>
- (Var ((P, 0), Logic.varifyT (fsT' --> T --> HOLogic.boolT)),
- Abs ("z", HOLogic.unitT, absfree (x, U, Q))))
- (pnames ~~ recTs ~~ rec_unique_frees ~~ rec_unique_concls) @
- map (fn (s, T) => (Var ((s, 0), Logic.varifyT T), Free (s, T)))
- rec_unique_frees)) induct_aux;
-
- fun obtain_fresh_name vs ths rec_fin_supp T (freshs1, freshs2, ctxt) =
- let
- val p = foldr1 HOLogic.mk_prod (vs @ freshs1);
- val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop
- (HOLogic.exists_const T $ Abs ("x", T,
- fresh_const T (fastype_of p) $ Bound 0 $ p)))
- (fn _ => EVERY
- [cut_facts_tac ths 1,
- REPEAT_DETERM (dresolve_tac (the (AList.lookup op = rec_fin_supp T)) 1),
- resolve_tac exists_fresh' 1,
- asm_simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]);
- val (([cx], ths), ctxt') = Obtain.result
- (fn _ => EVERY
- [etac exE 1,
- full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1,
- REPEAT (etac conjE 1)])
- [ex] ctxt
- in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end;
-
- val finite_ctxt_prems = map (fn aT =>
- HOLogic.mk_Trueprop
- (Const ("Finite_Set.finite", HOLogic.mk_setT aT --> HOLogic.boolT) $
- (Const ("Nominal.supp", fsT' --> HOLogic.mk_setT aT) $ rec_ctxt))) dt_atomTs;
-
- val rec_unique_thms = split_conj_thm (Goal.prove
- (ProofContext.init thy11) (map fst rec_unique_frees)
- (map (augment_sort thy11 fs_cp_sort)
- (List.concat finite_premss @ finite_ctxt_prems @ rec_prems @ rec_prems'))
- (augment_sort thy11 fs_cp_sort
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj rec_unique_concls)))
- (fn {prems, context} =>
- let
- val k = length rec_fns;
- val (finite_thss, ths1) = fold_map (fn T => fn xs =>
- apfst (pair T) (chop k xs)) dt_atomTs prems;
- val (finite_ctxt_ths, ths2) = chop (length dt_atomTs) ths1;
- val (P_ind_ths, fcbs) = chop k ths2;
- val P_ths = map (fn th => th RS mp) (split_conj_thm
- (Goal.prove context
- (map fst (rec_unique_frees'' @ rec_unique_frees')) []
- (augment_sort thy11 fs_cp_sort
- (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
- (map (fn (((x, y), S), P) => HOLogic.mk_imp
- (S $ Free x $ Free y, P $ (Free y)))
- (rec_unique_frees'' ~~ rec_unique_frees' ~~
- rec_sets ~~ rec_preds)))))
- (fn _ =>
- rtac rec_induct 1 THEN
- REPEAT ((resolve_tac P_ind_ths THEN_ALL_NEW assume_tac) 1))));
- val rec_fin_supp_thms' = map
- (fn (ths, (T, fin_ths)) => (T, map (curry op MRS fin_ths) ths))
- (rec_fin_supp_thms ~~ finite_thss);
- in EVERY
- ([rtac induct_aux_rec 1] @
- maps (fn ((_, finite_ths), finite_th) =>
- [cut_facts_tac (finite_th :: finite_ths) 1,
- asm_simp_tac (HOL_ss addsimps [supp_prod, finite_Un]) 1])
- (finite_thss ~~ finite_ctxt_ths) @
- maps (fn ((_, idxss), elim) => maps (fn idxs =>
- [full_simp_tac (HOL_ss addsimps [symmetric fresh_def, supp_prod, Un_iff]) 1,
- REPEAT_DETERM (eresolve_tac [conjE, ex1E] 1),
- rtac ex1I 1,
- (resolve_tac rec_intrs THEN_ALL_NEW atac) 1,
- rotate_tac ~1 1,
- ((DETERM o etac elim) THEN_ALL_NEW full_simp_tac
- (HOL_ss addsimps List.concat distinct_thms)) 1] @
- (if null idxs then [] else [hyp_subst_tac 1,
- SUBPROOF (fn {asms, concl, prems = prems', params, context = context', ...} =>
- let
- val SOME prem = find_first (can (HOLogic.dest_eq o
- HOLogic.dest_Trueprop o prop_of)) prems';
- val _ $ (_ $ lhs $ rhs) = prop_of prem;
- val _ $ (_ $ lhs' $ rhs') = term_of concl;
- val rT = fastype_of lhs';
- val (c, cargsl) = strip_comb lhs;
- val cargsl' = partition_cargs idxs cargsl;
- val boundsl = List.concat (map fst cargsl');
- val (_, cargsr) = strip_comb rhs;
- val cargsr' = partition_cargs idxs cargsr;
- val boundsr = List.concat (map fst cargsr');
- val (params1, _ :: params2) =
- chop (length params div 2) (map term_of params);
- val params' = params1 @ params2;
- val rec_prems = filter (fn th => case prop_of th of
- _ $ p => (case head_of p of
- Const (s, _) => s mem rec_set_names
- | _ => false)
- | _ => false) prems';
- val fresh_prems = filter (fn th => case prop_of th of
- _ $ (Const ("Nominal.fresh", _) $ _ $ _) => true
- | _ $ (Const ("Not", _) $ _) => true
- | _ => false) prems';
- val Ts = map fastype_of boundsl;
-
- val _ = warning "step 1: obtaining fresh names";
- val (freshs1, freshs2, context'') = fold
- (obtain_fresh_name (rec_ctxt :: rec_fns' @ params')
- (List.concat (map snd finite_thss) @
- finite_ctxt_ths @ rec_prems)
- rec_fin_supp_thms')
- Ts ([], [], context');
- val pi1 = map perm_of_pair (boundsl ~~ freshs1);
- val rpi1 = rev pi1;
- val pi2 = map perm_of_pair (boundsr ~~ freshs1);
- val rpi2 = rev pi2;
-
- val fresh_prems' = mk_not_sym fresh_prems;
- val freshs2' = mk_not_sym freshs2;
-
- (** as, bs, cs # K as ts, K bs us **)
- val _ = warning "step 2: as, bs, cs # K as ts, K bs us";
- val prove_fresh_ss = HOL_ss addsimps
- (finite_Diff :: List.concat fresh_thms @
- fs_atoms @ abs_fresh @ abs_supp @ fresh_atm);
- (* FIXME: avoid asm_full_simp_tac ? *)
- fun prove_fresh ths y x = Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (fresh_const
- (fastype_of x) (fastype_of y) $ x $ y))
- (fn _ => cut_facts_tac ths 1 THEN asm_full_simp_tac prove_fresh_ss 1);
- val constr_fresh_thms =
- map (prove_fresh fresh_prems lhs) boundsl @
- map (prove_fresh fresh_prems rhs) boundsr @
- map (prove_fresh freshs2 lhs) freshs1 @
- map (prove_fresh freshs2 rhs) freshs1;
-
- (** pi1 o (K as ts) = pi2 o (K bs us) **)
- val _ = warning "step 3: pi1 o (K as ts) = pi2 o (K bs us)";
- val pi1_pi2_eq = Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (fold_rev (mk_perm []) pi1 lhs, fold_rev (mk_perm []) pi2 rhs)))
- (fn _ => EVERY
- [cut_facts_tac constr_fresh_thms 1,
- asm_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh) 1,
- rtac prem 1]);
-
- (** pi1 o ts = pi2 o us **)
- val _ = warning "step 4: pi1 o ts = pi2 o us";
- val pi1_pi2_eqs = map (fn (t, u) =>
- Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (fold_rev (mk_perm []) pi1 t, fold_rev (mk_perm []) pi2 u)))
- (fn _ => EVERY
- [cut_facts_tac [pi1_pi2_eq] 1,
- asm_full_simp_tac (HOL_ss addsimps
- (calc_atm @ List.concat perm_simps' @
- fresh_prems' @ freshs2' @ abs_perm @
- alpha @ List.concat inject_thms)) 1]))
- (map snd cargsl' ~~ map snd cargsr');
-
- (** pi1^-1 o pi2 o us = ts **)
- val _ = warning "step 5: pi1^-1 o pi2 o us = ts";
- val rpi1_pi2_eqs = map (fn ((t, u), eq) =>
- Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (fold_rev (mk_perm []) (rpi1 @ pi2) u, t)))
- (fn _ => simp_tac (HOL_ss addsimps
- ((eq RS sym) :: perm_swap)) 1))
- (map snd cargsl' ~~ map snd cargsr' ~~ pi1_pi2_eqs);
-
- val (rec_prems1, rec_prems2) =
- chop (length rec_prems div 2) rec_prems;
-
- (** (ts, pi1^-1 o pi2 o vs) in rec_set **)
- val _ = warning "step 6: (ts, pi1^-1 o pi2 o vs) in rec_set";
- val rec_prems' = map (fn th =>
- let
- val _ $ (S $ x $ y) = prop_of th;
- val Const (s, _) = head_of S;
- val k = find_index (equal s) rec_set_names;
- val pi = rpi1 @ pi2;
- fun mk_pi z = fold_rev (mk_perm []) pi z;
- fun eqvt_tac p =
- let
- val U as Type (_, [Type (_, [T, _])]) = fastype_of p;
- val l = find_index (equal T) dt_atomTs;
- val th = List.nth (List.nth (rec_equiv_thms', l), k);
- val th' = Thm.instantiate ([],
- [(cterm_of thy11 (Var (("pi", 0), U)),
- cterm_of thy11 p)]) th;
- in rtac th' 1 end;
- val th' = Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (S $ mk_pi x $ mk_pi y))
- (fn _ => EVERY
- (map eqvt_tac pi @
- [simp_tac (HOL_ss addsimps (fresh_prems' @ freshs2' @
- perm_swap @ perm_fresh_fresh)) 1,
- rtac th 1]))
- in
- Simplifier.simplify
- (HOL_basic_ss addsimps rpi1_pi2_eqs) th'
- end) rec_prems2;
-
- val ihs = filter (fn th => case prop_of th of
- _ $ (Const ("All", _) $ _) => true | _ => false) prems';
-
- (** pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs **)
- val _ = warning "step 7: pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs";
- val rec_eqns = map (fn (th, ih) =>
- let
- val th' = th RS (ih RS spec RS mp) RS sym;
- val _ $ (_ $ lhs $ rhs) = prop_of th';
- fun strip_perm (_ $ _ $ t) = strip_perm t
- | strip_perm t = t;
- in
- Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (fold_rev (mk_perm []) pi1 lhs,
- fold_rev (mk_perm []) pi2 (strip_perm rhs))))
- (fn _ => simp_tac (HOL_basic_ss addsimps
- (th' :: perm_swap)) 1)
- end) (rec_prems' ~~ ihs);
-
- (** as # rs **)
- val _ = warning "step 8: as # rs";
- val rec_freshs = List.concat
- (map (fn (rec_prem, ih) =>
- let
- val _ $ (S $ x $ (y as Free (_, T))) =
- prop_of rec_prem;
- val k = find_index (equal S) rec_sets;
- val atoms = List.concat (List.mapPartial (fn (bs, z) =>
- if z = x then NONE else SOME bs) cargsl')
- in
- map (fn a as Free (_, aT) =>
- let val l = find_index (equal aT) dt_atomTs;
- in
- Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (fresh_const aT T $ a $ y))
- (fn _ => EVERY
- (rtac (List.nth (List.nth (rec_fresh_thms, l), k)) 1 ::
- map (fn th => rtac th 1)
- (snd (List.nth (finite_thss, l))) @
- [rtac rec_prem 1, rtac ih 1,
- REPEAT_DETERM (resolve_tac fresh_prems 1)]))
- end) atoms
- end) (rec_prems1 ~~ ihs));
-
- (** as # fK as ts rs , bs # fK bs us vs **)
- val _ = warning "step 9: as # fK as ts rs , bs # fK bs us vs";
- fun prove_fresh_result (a as Free (_, aT)) =
- Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ rhs'))
- (fn _ => EVERY
- [resolve_tac fcbs 1,
- REPEAT_DETERM (resolve_tac
- (fresh_prems @ rec_freshs) 1),
- REPEAT_DETERM (resolve_tac (maps snd rec_fin_supp_thms') 1
- THEN resolve_tac rec_prems 1),
- resolve_tac P_ind_ths 1,
- REPEAT_DETERM (resolve_tac (P_ths @ rec_prems) 1)]);
-
- val fresh_results'' = map prove_fresh_result boundsl;
-
- fun prove_fresh_result'' ((a as Free (_, aT), b), th) =
- let val th' = Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (fresh_const aT rT $
- fold_rev (mk_perm []) (rpi2 @ pi1) a $
- fold_rev (mk_perm []) (rpi2 @ pi1) rhs'))
- (fn _ => simp_tac (HOL_ss addsimps fresh_bij) 1 THEN
- rtac th 1)
- in
- Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (fresh_const aT rT $ b $ lhs'))
- (fn _ => EVERY
- [cut_facts_tac [th'] 1,
- full_simp_tac (Simplifier.theory_context thy11 HOL_ss
- addsimps rec_eqns @ pi1_pi2_eqs @ perm_swap
- addsimprocs [NominalPermeq.perm_simproc_app]) 1,
- full_simp_tac (HOL_ss addsimps (calc_atm @
- fresh_prems' @ freshs2' @ perm_fresh_fresh)) 1])
- end;
-
- val fresh_results = fresh_results'' @ map prove_fresh_result''
- (boundsl ~~ boundsr ~~ fresh_results'');
-
- (** cs # fK as ts rs , cs # fK bs us vs **)
- val _ = warning "step 10: cs # fK as ts rs , cs # fK bs us vs";
- fun prove_fresh_result' recs t (a as Free (_, aT)) =
- Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ t))
- (fn _ => EVERY
- [cut_facts_tac recs 1,
- REPEAT_DETERM (dresolve_tac
- (the (AList.lookup op = rec_fin_supp_thms' aT)) 1),
- NominalPermeq.fresh_guess_tac
- (HOL_ss addsimps (freshs2 @
- fs_atoms @ fresh_atm @
- List.concat (map snd finite_thss))) 1]);
-
- val fresh_results' =
- map (prove_fresh_result' rec_prems1 rhs') freshs1 @
- map (prove_fresh_result' rec_prems2 lhs') freshs1;
-
- (** pi1 o (fK as ts rs) = pi2 o (fK bs us vs) **)
- val _ = warning "step 11: pi1 o (fK as ts rs) = pi2 o (fK bs us vs)";
- val pi1_pi2_result = Goal.prove context'' [] []
- (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (fold_rev (mk_perm []) pi1 rhs', fold_rev (mk_perm []) pi2 lhs')))
- (fn _ => simp_tac (Simplifier.context context'' HOL_ss
- addsimps pi1_pi2_eqs @ rec_eqns
- addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN
- TRY (simp_tac (HOL_ss addsimps
- (fresh_prems' @ freshs2' @ calc_atm @ perm_fresh_fresh)) 1));
-
- val _ = warning "final result";
- val final = Goal.prove context'' [] [] (term_of concl)
- (fn _ => cut_facts_tac [pi1_pi2_result RS sym] 1 THEN
- full_simp_tac (HOL_basic_ss addsimps perm_fresh_fresh @
- fresh_results @ fresh_results') 1);
- val final' = ProofContext.export context'' context' [final];
- val _ = warning "finished!"
- in
- resolve_tac final' 1
- end) context 1])) idxss) (ndescr ~~ rec_elims))
- end));
-
- val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
-
- (* define primrec combinators *)
-
- val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
- val reccomb_names = map (Sign.full_bname thy11)
- (if length descr'' = 1 then [big_reccomb_name] else
- (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
- (1 upto (length descr''))));
- val reccombs = map (fn ((name, T), T') => list_comb
- (Const (name, rec_fn_Ts @ [T] ---> T'), rec_fns))
- (reccomb_names ~~ recTs ~~ rec_result_Ts);
-
- val (reccomb_defs, thy12) =
- thy11
- |> Sign.add_consts_i (map (fn ((name, T), T') =>
- (Binding.name (Long_Name.base_name name), rec_fn_Ts @ [T] ---> T', NoSyn))
- (reccomb_names ~~ recTs ~~ rec_result_Ts))
- |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
- (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T,
- Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
- set $ Free ("x", T) $ Free ("y", T'))))))
- (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts));
-
- (* prove characteristic equations for primrec combinators *)
-
- val rec_thms = map (fn (prems, concl) =>
- let
- val _ $ (_ $ (_ $ x) $ _) = concl;
- val (_, cargs) = strip_comb x;
- val ps = map (fn (x as Free (_, T), i) =>
- (Free ("x" ^ string_of_int i, T), x)) (cargs ~~ (1 upto length cargs));
- val concl' = subst_atomic_types (rec_result_Ts' ~~ rec_result_Ts) concl;
- val prems' = List.concat finite_premss @ finite_ctxt_prems @
- rec_prems @ rec_prems' @ map (subst_atomic ps) prems;
- fun solve rules prems = resolve_tac rules THEN_ALL_NEW
- (resolve_tac prems THEN_ALL_NEW atac)
- in
- Goal.prove_global thy12 []
- (map (augment_sort thy12 fs_cp_sort) prems')
- (augment_sort thy12 fs_cp_sort concl')
- (fn {prems, ...} => EVERY
- [rewrite_goals_tac reccomb_defs,
- rtac the1_equality 1,
- solve rec_unique_thms prems 1,
- resolve_tac rec_intrs 1,
- REPEAT (solve (prems @ rec_total_thms) prems 1)])
- end) (rec_eq_prems ~~
- DatatypeProp.make_primrecs new_type_names descr' sorts thy12);
-
- val dt_infos = map (make_dt_info pdescr sorts induct reccomb_names rec_thms)
- ((0 upto length descr1 - 1) ~~ descr1 ~~ distinct_thms ~~ inject_thms);
-
- (* FIXME: theorems are stored in database for testing only *)
- val (_, thy13) = thy12 |>
- PureThy.add_thmss
- [((Binding.name "rec_equiv", List.concat rec_equiv_thms), []),
- ((Binding.name "rec_equiv'", List.concat rec_equiv_thms'), []),
- ((Binding.name "rec_fin_supp", List.concat rec_fin_supp_thms), []),
- ((Binding.name "rec_fresh", List.concat rec_fresh_thms), []),
- ((Binding.name "rec_unique", map standard rec_unique_thms), []),
- ((Binding.name "recs", rec_thms), [])] ||>
- Sign.parent_path ||>
- map_nominal_datatypes (fold Symtab.update dt_infos);
-
- in
- thy13
- end;
-
-val add_nominal_datatype = gen_add_nominal_datatype DatatypePackage.read_typ;
-
-
-(* FIXME: The following stuff should be exported by DatatypePackage *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val datatype_decl =
- Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix --
- (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix));
-
-fun mk_datatype args =
- let
- val names = map (fn ((((NONE, _), t), _), _) => t | ((((SOME t, _), _), _), _) => t) args;
- val specs = map (fn ((((_, vs), t), mx), cons) =>
- (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
- in add_nominal_datatype DatatypeAux.default_datatype_config names specs end;
-
-val _ =
- OuterSyntax.command "nominal_datatype" "define inductive datatypes" K.thy_decl
- (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
-
-end;
-
-end
--- a/src/HOL/Nominal/nominal_primrec.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Nominal/nominal_primrec.ML Fri Jun 19 17:23:21 2009 +0200
@@ -3,7 +3,7 @@
Author: Stefan Berghofer, TU Muenchen
Package for defining functions on nominal datatypes by primitive recursion.
-Taken from HOL/Tools/primrec_package.ML
+Taken from HOL/Tools/primrec.ML
*)
signature NOMINAL_PRIMREC =
@@ -223,7 +223,7 @@
(* find datatypes which contain all datatypes in tnames' *)
-fun find_dts (dt_info : NominalPackage.nominal_datatype_info Symtab.table) _ [] = []
+fun find_dts (dt_info : Nominal.nominal_datatype_info Symtab.table) _ [] = []
| find_dts dt_info tnames' (tname::tnames) =
(case Symtab.lookup dt_info tname of
NONE => primrec_err (quote tname ^ " is not a nominal datatype")
@@ -247,7 +247,7 @@
val eqns' = map unquantify spec'
val eqns = fold_rev (process_eqn lthy (fn v => Variable.is_fixed lthy v
orelse exists (fn ((w, _), _) => v = Binding.name_of w) fixes)) spec' [];
- val dt_info = NominalPackage.get_nominal_datatypes (ProofContext.theory_of lthy);
+ val dt_info = Nominal.get_nominal_datatypes (ProofContext.theory_of lthy);
val lsrs :: lsrss = maps (fn (_, (_, _, eqns)) =>
map (fn (_, (ls, _, rs, _, _)) => ls @ rs) eqns) eqns
val _ =
--- a/src/HOL/Product_Type.thy Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Product_Type.thy Fri Jun 19 17:23:21 2009 +0200
@@ -9,7 +9,7 @@
imports Inductive
uses
("Tools/split_rule.ML")
- ("Tools/inductive_set_package.ML")
+ ("Tools/inductive_set.ML")
("Tools/inductive_realizer.ML")
("Tools/datatype_package/datatype_realizer.ML")
begin
@@ -1151,8 +1151,8 @@
use "Tools/inductive_realizer.ML"
setup InductiveRealizer.setup
-use "Tools/inductive_set_package.ML"
-setup InductiveSetPackage.setup
+use "Tools/inductive_set.ML"
+setup Inductive_Set.setup
use "Tools/datatype_package/datatype_realizer.ML"
setup DatatypeRealizer.setup
--- a/src/HOL/Recdef.thy Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Recdef.thy Fri Jun 19 17:23:21 2009 +0200
@@ -16,7 +16,7 @@
("Tools/TFL/thry.ML")
("Tools/TFL/tfl.ML")
("Tools/TFL/post.ML")
- ("Tools/recdef_package.ML")
+ ("Tools/recdef.ML")
begin
text{** This form avoids giant explosions in proofs. NOTE USE OF ==*}
@@ -76,8 +76,8 @@
use "Tools/TFL/thry.ML"
use "Tools/TFL/tfl.ML"
use "Tools/TFL/post.ML"
-use "Tools/recdef_package.ML"
-setup RecdefPackage.setup
+use "Tools/recdef.ML"
+setup Recdef.setup
lemmas [recdef_simp] =
inv_image_def
--- a/src/HOL/Record.thy Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Record.thy Fri Jun 19 17:23:21 2009 +0200
@@ -1,5 +1,4 @@
(* Title: HOL/Record.thy
- ID: $Id$
Author: Wolfgang Naraschewski, Norbert Schirmer and Markus Wenzel, TU Muenchen
*)
@@ -7,7 +6,7 @@
theory Record
imports Product_Type
-uses ("Tools/record_package.ML")
+uses ("Tools/record.ML")
begin
lemma prop_subst: "s = t \<Longrightarrow> PROP P t \<Longrightarrow> PROP P s"
@@ -56,7 +55,7 @@
"_record_scheme" :: "[fields, 'a] => 'a" ("(3\<lparr>_,/ (2\<dots> =/ _)\<rparr>)")
"_record_update" :: "['a, updates] => 'b" ("_/(3\<lparr>_\<rparr>)" [900,0] 900)
-use "Tools/record_package.ML"
-setup RecordPackage.setup
+use "Tools/record.ML"
+setup Record.setup
end
--- a/src/HOL/Statespace/state_fun.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Statespace/state_fun.ML Fri Jun 19 17:23:21 2009 +0200
@@ -74,7 +74,7 @@
val string_eq_simp_tac =
simp_tac (HOL_basic_ss
addsimps (thms "list.inject"@thms "char.inject"@simp_thms)
- addsimprocs [DatatypePackage.distinct_simproc,lazy_conj_simproc]
+ addsimprocs [Datatype.distinct_simproc,lazy_conj_simproc]
addcongs [thm "block_conj_cong"])
end;
@@ -89,7 +89,7 @@
in
val lookup_ss = (HOL_basic_ss
addsimps (thms "list.inject"@thms "char.inject"@simp_thms@rules)
- addsimprocs [DatatypePackage.distinct_simproc,lazy_conj_simproc]
+ addsimprocs [Datatype.distinct_simproc,lazy_conj_simproc]
addcongs [thm "block_conj_cong"]
addSolver StateSpace.distinctNameSolver)
end;
@@ -167,7 +167,7 @@
val meta_ext = thm "StateFun.meta_ext";
val o_apply = thm "Fun.o_apply";
val ss' = (HOL_ss addsimps (update_apply::o_apply::thms "list.inject"@thms "char.inject")
- addsimprocs [DatatypePackage.distinct_simproc,lazy_conj_simproc,StateSpace.distinct_simproc]
+ addsimprocs [Datatype.distinct_simproc,lazy_conj_simproc,StateSpace.distinct_simproc]
addcongs [thm "block_conj_cong"]);
in
val update_simproc =
@@ -267,7 +267,7 @@
val swap_ex_eq = thm "StateFun.swap_ex_eq";
fun is_selector thy T sel =
let
- val (flds,more) = RecordPackage.get_recT_fields thy T
+ val (flds,more) = Record.get_recT_fields thy T
in member (fn (s,(n,_)) => n=s) (more::flds) sel
end;
in
@@ -340,7 +340,7 @@
| mkName (TFree (x,_)) = mkUpper (Long_Name.base_name x)
| mkName (TVar ((x,_),_)) = mkUpper (Long_Name.base_name x);
-fun is_datatype thy n = is_some (Symtab.lookup (DatatypePackage.get_datatypes thy) n);
+fun is_datatype thy n = is_some (Symtab.lookup (Datatype.get_datatypes thy) n);
fun mk_map "List.list" = Syntax.const "List.map"
| mk_map n = Syntax.const ("StateFun.map_" ^ Long_Name.base_name n);
--- a/src/HOL/Statespace/state_space.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Statespace/state_space.ML Fri Jun 19 17:23:21 2009 +0200
@@ -585,8 +585,8 @@
end
handle ERROR msg => cat_error msg ("Failed to define statespace " ^ quote name);
-val define_statespace = gen_define_statespace RecordPackage.read_typ NONE;
-val define_statespace_i = gen_define_statespace RecordPackage.cert_typ;
+val define_statespace = gen_define_statespace Record.read_typ NONE;
+val define_statespace_i = gen_define_statespace Record.cert_typ;
(*** parse/print - translations ***)
--- a/src/HOL/Tools/TFL/casesplit.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Tools/TFL/casesplit.ML Fri Jun 19 17:23:21 2009 +0200
@@ -90,7 +90,7 @@
(* get the case_thm (my version) from a type *)
fun case_thm_of_ty sgn ty =
let
- val dtypestab = DatatypePackage.get_datatypes sgn;
+ val dtypestab = Datatype.get_datatypes sgn;
val ty_str = case ty of
Type(ty_str, _) => ty_str
| TFree(s,_) => error ("Free type: " ^ s)
--- a/src/HOL/Tools/TFL/tfl.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Tools/TFL/tfl.ML Fri Jun 19 17:23:21 2009 +0200
@@ -446,7 +446,7 @@
slow.*)
val case_ss = Simplifier.theory_context theory
(HOL_basic_ss addcongs
- (map (#weak_case_cong o snd) o Symtab.dest o DatatypePackage.get_datatypes) theory addsimps case_rewrites)
+ (map (#weak_case_cong o snd) o Symtab.dest o Datatype.get_datatypes) theory addsimps case_rewrites)
val corollaries' = map (Simplifier.simplify case_ss) corollaries
val extract = R.CONTEXT_REWRITE_RULE
(f, [R], @{thm cut_apply}, meta_tflCongs@context_congs)
--- a/src/HOL/Tools/TFL/thry.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Tools/TFL/thry.ML Fri Jun 19 17:23:21 2009 +0200
@@ -60,20 +60,20 @@
*---------------------------------------------------------------------------*)
fun match_info thy dtco =
- case (DatatypePackage.get_datatype thy dtco,
- DatatypePackage.get_datatype_constrs thy dtco) of
+ case (Datatype.get_datatype thy dtco,
+ Datatype.get_datatype_constrs thy dtco) of
(SOME { case_name, ... }, SOME constructors) =>
SOME {case_const = Const (case_name, Sign.the_const_type thy case_name), constructors = map Const constructors}
| _ => NONE;
-fun induct_info thy dtco = case DatatypePackage.get_datatype thy dtco of
+fun induct_info thy dtco = case Datatype.get_datatype thy dtco of
NONE => NONE
| SOME {nchotomy, ...} =>
SOME {nchotomy = nchotomy,
- constructors = (map Const o the o DatatypePackage.get_datatype_constrs thy) dtco};
+ constructors = (map Const o the o Datatype.get_datatype_constrs thy) dtco};
fun extract_info thy =
- let val infos = (map snd o Symtab.dest o DatatypePackage.get_datatypes) thy
+ let val infos = (map snd o Symtab.dest o Datatype.get_datatypes) thy
in {case_congs = map (mk_meta_eq o #case_cong) infos,
case_rewrites = List.concat (map (map mk_meta_eq o #case_rewrites) infos)}
end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/choice_specification.ML Fri Jun 19 17:23:21 2009 +0200
@@ -0,0 +1,257 @@
+(* Title: HOL/Tools/choice_specification.ML
+ Author: Sebastian Skalberg, TU Muenchen
+
+Package for defining constants by specification.
+*)
+
+signature CHOICE_SPECIFICATION =
+sig
+ val add_specification: string option -> (bstring * xstring * bool) list ->
+ theory * thm -> theory * thm
+end
+
+structure Choice_Specification: CHOICE_SPECIFICATION =
+struct
+
+(* actual code *)
+
+local
+ fun mk_definitional [] arg = arg
+ | mk_definitional ((thname,cname,covld)::cos) (thy,thm) =
+ case HOLogic.dest_Trueprop (concl_of thm) of
+ Const("Ex",_) $ P =>
+ let
+ val ctype = domain_type (type_of P)
+ val cname_full = Sign.intern_const thy cname
+ val cdefname = if thname = ""
+ then Thm.def_name (Long_Name.base_name cname)
+ else thname
+ val def_eq = Logic.mk_equals (Const(cname_full,ctype),
+ HOLogic.choice_const ctype $ P)
+ val (thms, thy') = PureThy.add_defs covld [((Binding.name cdefname, def_eq),[])] thy
+ val thm' = [thm,hd thms] MRS @{thm exE_some}
+ in
+ mk_definitional cos (thy',thm')
+ end
+ | _ => raise THM ("Internal error: Bad specification theorem",0,[thm])
+
+ fun mk_axiomatic axname cos arg =
+ let
+ fun process [] (thy,tm) =
+ let
+ val (thms, thy') = PureThy.add_axioms [((Binding.name axname, HOLogic.mk_Trueprop tm),[])] thy
+ in
+ (thy',hd thms)
+ end
+ | process ((thname,cname,covld)::cos) (thy,tm) =
+ case tm of
+ Const("Ex",_) $ P =>
+ let
+ val ctype = domain_type (type_of P)
+ val cname_full = Sign.intern_const thy cname
+ val cdefname = if thname = ""
+ then Thm.def_name (Long_Name.base_name cname)
+ else thname
+ val co = Const(cname_full,ctype)
+ val thy' = Theory.add_finals_i covld [co] thy
+ val tm' = case P of
+ Abs(_, _, bodt) => subst_bound (co, bodt)
+ | _ => P $ co
+ in
+ process cos (thy',tm')
+ end
+ | _ => raise TERM ("Internal error: Bad specification theorem",[tm])
+ in
+ process cos arg
+ end
+
+in
+fun proc_exprop axiomatic cos arg =
+ case axiomatic of
+ SOME axname => mk_axiomatic axname cos (apsnd (HOLogic.dest_Trueprop o concl_of) arg)
+ | NONE => mk_definitional cos arg
+end
+
+fun add_specification axiomatic cos arg =
+ arg |> apsnd Thm.freezeT
+ |> proc_exprop axiomatic cos
+ |> apsnd standard
+
+
+(* Collect all intances of constants in term *)
+
+fun collect_consts ( t $ u,tms) = collect_consts (u,collect_consts (t,tms))
+ | collect_consts ( Abs(_,_,t),tms) = collect_consts (t,tms)
+ | collect_consts (tm as Const _,tms) = insert (op aconv) tm tms
+ | collect_consts ( _,tms) = tms
+
+(* Complementing Type.varify... *)
+
+fun unvarify t fmap =
+ let
+ val fmap' = map Library.swap fmap
+ fun unthaw (f as (a, S)) =
+ (case AList.lookup (op =) fmap' a of
+ NONE => TVar f
+ | SOME (b, _) => TFree (b, S))
+ in
+ map_types (map_type_tvar unthaw) t
+ end
+
+(* The syntactic meddling needed to setup add_specification for work *)
+
+fun process_spec axiomatic cos alt_props thy =
+ let
+ fun zip3 [] [] [] = []
+ | zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs
+ | zip3 _ _ _ = error "Choice_Specification.process_spec internal error"
+
+ fun myfoldr f [x] = x
+ | myfoldr f (x::xs) = f (x,myfoldr f xs)
+ | myfoldr f [] = error "Choice_Specification.process_spec internal error"
+
+ val rew_imps = alt_props |>
+ map (ObjectLogic.atomize o Thm.cterm_of thy o Syntax.read_prop_global thy o snd)
+ val props' = rew_imps |>
+ map (HOLogic.dest_Trueprop o term_of o snd o Thm.dest_equals o cprop_of)
+
+ fun proc_single prop =
+ let
+ val frees = OldTerm.term_frees prop
+ val _ = forall (fn v => Sign.of_sort thy (type_of v,HOLogic.typeS)) frees
+ orelse error "Specificaton: Only free variables of sort 'type' allowed"
+ val prop_closed = List.foldr (fn ((vname,T),prop) => HOLogic.mk_all (vname,T,prop)) prop (map dest_Free frees)
+ in
+ (prop_closed,frees)
+ end
+
+ val props'' = map proc_single props'
+ val frees = map snd props''
+ val prop = myfoldr HOLogic.mk_conj (map fst props'')
+ val cprop = cterm_of thy (HOLogic.mk_Trueprop prop)
+
+ val (vmap, prop_thawed) = Type.varify [] prop
+ val thawed_prop_consts = collect_consts (prop_thawed,[])
+ val (altcos,overloaded) = Library.split_list cos
+ val (names,sconsts) = Library.split_list altcos
+ val consts = map (Syntax.read_term_global thy) sconsts
+ val _ = not (Library.exists (not o Term.is_Const) consts)
+ orelse error "Specification: Non-constant found as parameter"
+
+ fun proc_const c =
+ let
+ val (_, c') = Type.varify [] c
+ val (cname,ctyp) = dest_Const c'
+ in
+ case List.filter (fn t => let val (name,typ) = dest_Const t
+ in name = cname andalso Sign.typ_equiv thy (typ, ctyp)
+ end) thawed_prop_consts of
+ [] => error ("Specification: No suitable instances of constant \"" ^ Syntax.string_of_term_global thy c ^ "\" found")
+ | [cf] => unvarify cf vmap
+ | _ => error ("Specification: Several variations of \"" ^ Syntax.string_of_term_global thy c ^ "\" found (try applying explicit type constraints)")
+ end
+ val proc_consts = map proc_const consts
+ fun mk_exist (c,prop) =
+ let
+ val T = type_of c
+ val cname = Long_Name.base_name (fst (dest_Const c))
+ val vname = if Syntax.is_identifier cname
+ then cname
+ else "x"
+ in
+ HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c,prop))
+ end
+ val ex_prop = List.foldr mk_exist prop proc_consts
+ val cnames = map (fst o dest_Const) proc_consts
+ fun post_process (arg as (thy,thm)) =
+ let
+ fun inst_all thy (thm,v) =
+ let
+ val cv = cterm_of thy v
+ val cT = ctyp_of_term cv
+ val spec' = instantiate' [SOME cT] [NONE,SOME cv] spec
+ in
+ thm RS spec'
+ end
+ fun remove_alls frees thm =
+ Library.foldl (inst_all (Thm.theory_of_thm thm)) (thm,frees)
+ fun process_single ((name,atts),rew_imp,frees) args =
+ let
+ fun undo_imps thm =
+ equal_elim (symmetric rew_imp) thm
+
+ fun add_final (arg as (thy, thm)) =
+ if name = ""
+ then arg |> Library.swap
+ else (writeln (" " ^ name ^ ": " ^ (Display.string_of_thm thm));
+ PureThy.store_thm (Binding.name name, thm) thy)
+ in
+ args |> apsnd (remove_alls frees)
+ |> apsnd undo_imps
+ |> apsnd standard
+ |> Thm.theory_attributes (map (Attrib.attribute thy) atts)
+ |> add_final
+ |> Library.swap
+ end
+
+ fun process_all [proc_arg] args =
+ process_single proc_arg args
+ | process_all (proc_arg::rest) (thy,thm) =
+ let
+ val single_th = thm RS conjunct1
+ val rest_th = thm RS conjunct2
+ val (thy',_) = process_single proc_arg (thy,single_th)
+ in
+ process_all rest (thy',rest_th)
+ end
+ | process_all [] _ = error "Choice_Specification.process_spec internal error"
+ val alt_names = map fst alt_props
+ val _ = if exists (fn(name,_) => not (name = "")) alt_names
+ then writeln "specification"
+ else ()
+ in
+ arg |> apsnd Thm.freezeT
+ |> process_all (zip3 alt_names rew_imps frees)
+ end
+
+ fun after_qed [[thm]] = ProofContext.theory (fn thy =>
+ #1 (post_process (add_specification axiomatic (zip3 names cnames overloaded) (thy, thm))));
+ in
+ thy
+ |> ProofContext.init
+ |> Proof.theorem_i NONE after_qed [[(HOLogic.mk_Trueprop ex_prop, [])]]
+ end;
+
+
+(* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val opt_name = Scan.optional (P.name --| P.$$$ ":") ""
+val opt_overloaded = P.opt_keyword "overloaded";
+
+val specification_decl =
+ P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
+ Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.prop)
+
+val _ =
+ OuterSyntax.command "specification" "define constants by specification" K.thy_goal
+ (specification_decl >> (fn (cos,alt_props) =>
+ Toplevel.print o (Toplevel.theory_to_proof
+ (process_spec NONE cos alt_props))))
+
+val ax_specification_decl =
+ P.name --
+ (P.$$$ "(" |-- Scan.repeat1 (opt_name -- P.term -- opt_overloaded) --| P.$$$ ")" --
+ Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.prop))
+
+val _ =
+ OuterSyntax.command "ax_specification" "define constants by specification" K.thy_goal
+ (ax_specification_decl >> (fn (axname,(cos,alt_props)) =>
+ Toplevel.print o (Toplevel.theory_to_proof
+ (process_spec (SOME axname) cos alt_props))))
+
+end
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/datatype_package/datatype.ML Fri Jun 19 17:23:21 2009 +0200
@@ -0,0 +1,705 @@
+(* Title: HOL/Tools/datatype.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Datatype package for Isabelle/HOL.
+*)
+
+signature DATATYPE =
+sig
+ type datatype_config = DatatypeAux.datatype_config
+ type datatype_info = DatatypeAux.datatype_info
+ type descr = DatatypeAux.descr
+ val get_datatypes : theory -> datatype_info Symtab.table
+ val get_datatype : theory -> string -> datatype_info option
+ val the_datatype : theory -> string -> datatype_info
+ val datatype_of_constr : theory -> string -> datatype_info option
+ val datatype_of_case : theory -> string -> datatype_info option
+ val the_datatype_spec : theory -> string -> (string * sort) list * (string * typ list) list
+ val the_datatype_descr : theory -> string list
+ -> descr * (string * sort) list * string list
+ * (string list * string list) * (typ list * typ list)
+ val get_datatype_constrs : theory -> string -> (string * typ) list option
+ val distinct_simproc : simproc
+ val make_case : Proof.context -> bool -> string list -> term ->
+ (term * term) list -> term * (term * (int * bool)) list
+ val strip_case : Proof.context -> bool -> term -> (term * (term * term) list) option
+ val read_typ: theory ->
+ (typ list * (string * sort) list) * string -> typ list * (string * sort) list
+ val interpretation : (datatype_config -> string list -> theory -> theory) -> theory -> theory
+ type rules = {distinct : thm list list,
+ inject : thm list list,
+ exhaustion : thm list,
+ rec_thms : thm list,
+ case_thms : thm list list,
+ split_thms : (thm * thm) list,
+ induction : thm,
+ simps : thm list}
+ val rep_datatype : datatype_config -> (rules -> Proof.context -> Proof.context)
+ -> string list option -> term list -> theory -> Proof.state;
+ val rep_datatype_cmd : string list option -> string list -> theory -> Proof.state
+ val add_datatype : datatype_config -> string list -> (string list * binding * mixfix *
+ (binding * typ list * mixfix) list) list -> theory -> rules * theory
+ val add_datatype_cmd : string list -> (string list * binding * mixfix *
+ (binding * string list * mixfix) list) list -> theory -> rules * theory
+ val setup: theory -> theory
+ val print_datatypes : theory -> unit
+end;
+
+structure Datatype : DATATYPE =
+struct
+
+open DatatypeAux;
+
+
+(* theory data *)
+
+structure DatatypesData = TheoryDataFun
+(
+ type T =
+ {types: datatype_info Symtab.table,
+ constrs: datatype_info Symtab.table,
+ cases: datatype_info Symtab.table};
+
+ val empty =
+ {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty};
+ val copy = I;
+ val extend = I;
+ fun merge _
+ ({types = types1, constrs = constrs1, cases = cases1},
+ {types = types2, constrs = constrs2, cases = cases2}) =
+ {types = Symtab.merge (K true) (types1, types2),
+ constrs = Symtab.merge (K true) (constrs1, constrs2),
+ cases = Symtab.merge (K true) (cases1, cases2)};
+);
+
+val get_datatypes = #types o DatatypesData.get;
+val map_datatypes = DatatypesData.map;
+
+fun print_datatypes thy =
+ Pretty.writeln (Pretty.strs ("datatypes:" ::
+ map #1 (NameSpace.extern_table (Sign.type_space thy, get_datatypes thy))));
+
+
+(** theory information about datatypes **)
+
+fun put_dt_infos (dt_infos : (string * datatype_info) list) =
+ map_datatypes (fn {types, constrs, cases} =>
+ {types = fold Symtab.update dt_infos types,
+ constrs = fold Symtab.default (*conservative wrt. overloaded constructors*)
+ (maps (fn (_, info as {descr, index, ...}) => map (rpair info o fst)
+ (#3 (the (AList.lookup op = descr index)))) dt_infos) constrs,
+ cases = fold Symtab.update
+ (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)
+ cases});
+
+val get_datatype = Symtab.lookup o get_datatypes;
+
+fun the_datatype thy name = (case get_datatype thy name of
+ SOME info => info
+ | NONE => error ("Unknown datatype " ^ quote name));
+
+val datatype_of_constr = Symtab.lookup o #constrs o DatatypesData.get;
+val datatype_of_case = Symtab.lookup o #cases o DatatypesData.get;
+
+fun get_datatype_descr thy dtco =
+ get_datatype thy dtco
+ |> Option.map (fn info as { descr, index, ... } =>
+ (info, (((fn SOME (_, dtys, cos) => (dtys, cos)) o AList.lookup (op =) descr) index)));
+
+fun the_datatype_spec thy dtco =
+ let
+ val info as { descr, index, sorts = raw_sorts, ... } = the_datatype thy dtco;
+ val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr index;
+ val sorts = map ((fn v => (v, (the o AList.lookup (op =) raw_sorts) v))
+ o DatatypeAux.dest_DtTFree) dtys;
+ val cos = map
+ (fn (co, tys) => (co, map (DatatypeAux.typ_of_dtyp descr sorts) tys)) raw_cos;
+ in (sorts, cos) end;
+
+fun the_datatype_descr thy (raw_tycos as raw_tyco :: _) =
+ let
+ val info = the_datatype thy raw_tyco;
+ val descr = #descr info;
+
+ val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr (#index info);
+ val vs = map ((fn v => (v, (the o AList.lookup (op =) (#sorts info)) v))
+ o dest_DtTFree) dtys;
+
+ fun is_DtTFree (DtTFree _) = true
+ | is_DtTFree _ = false
+ val k = find_index (fn (_, (_, dTs, _)) => not (forall is_DtTFree dTs)) descr;
+ val protoTs as (dataTs, _) = chop k descr
+ |> (pairself o map) (fn (_, (tyco, dTs, _)) => (tyco, map (typ_of_dtyp descr vs) dTs));
+
+ val tycos = map fst dataTs;
+ val _ = if gen_eq_set (op =) (tycos, raw_tycos) then ()
+ else error ("Type constructors " ^ commas (map quote raw_tycos)
+ ^ "do not belong exhaustively to one mutual recursive datatype");
+
+ val (Ts, Us) = (pairself o map) Type protoTs;
+
+ val names = map Long_Name.base_name (the_default tycos (#alt_names info));
+ val (auxnames, _) = Name.make_context names
+ |> fold_map (yield_singleton Name.variants o name_of_typ) Us
+
+ in (descr, vs, tycos, (names, auxnames), (Ts, Us)) end;
+
+fun get_datatype_constrs thy dtco =
+ case try (the_datatype_spec thy) dtco
+ of SOME (sorts, cos) =>
+ let
+ fun subst (v, sort) = TVar ((v, 0), sort);
+ fun subst_ty (TFree v) = subst v
+ | subst_ty ty = ty;
+ val dty = Type (dtco, map subst sorts);
+ fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty);
+ in SOME (map mk_co cos) end
+ | NONE => NONE;
+
+
+(** induct method setup **)
+
+(* case names *)
+
+local
+
+fun dt_recs (DtTFree _) = []
+ | dt_recs (DtType (_, dts)) = maps dt_recs dts
+ | dt_recs (DtRec i) = [i];
+
+fun dt_cases (descr: descr) (_, args, constrs) =
+ let
+ fun the_bname i = Long_Name.base_name (#1 (the (AList.lookup (op =) descr i)));
+ val bnames = map the_bname (distinct (op =) (maps dt_recs args));
+ in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end;
+
+
+fun induct_cases descr =
+ DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr));
+
+fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i));
+
+in
+
+fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
+
+fun mk_case_names_exhausts descr new =
+ map (RuleCases.case_names o exhaust_cases descr o #1)
+ (filter (fn ((_, (name, _, _))) => member (op =) new name) descr);
+
+end;
+
+fun add_rules simps case_thms rec_thms inject distinct
+ weak_case_congs cong_att =
+ PureThy.add_thmss [((Binding.name "simps", simps), []),
+ ((Binding.empty, flat case_thms @
+ flat distinct @ rec_thms), [Simplifier.simp_add]),
+ ((Binding.empty, rec_thms), [Code.add_default_eqn_attribute]),
+ ((Binding.empty, flat inject), [iff_add]),
+ ((Binding.empty, map (fn th => th RS notE) (flat distinct)), [Classical.safe_elim NONE]),
+ ((Binding.empty, weak_case_congs), [cong_att])]
+ #> snd;
+
+
+(* add_cases_induct *)
+
+fun add_cases_induct infos induction thy =
+ let
+ val inducts = ProjectRule.projections (ProofContext.init thy) induction;
+
+ fun named_rules (name, {index, exhaustion, ...}: datatype_info) =
+ [((Binding.empty, nth inducts index), [Induct.induct_type name]),
+ ((Binding.empty, exhaustion), [Induct.cases_type name])];
+ fun unnamed_rule i =
+ ((Binding.empty, nth inducts i), [Thm.kind_internal, Induct.induct_type ""]);
+ in
+ thy |> PureThy.add_thms
+ (maps named_rules infos @
+ map unnamed_rule (length infos upto length inducts - 1)) |> snd
+ |> PureThy.add_thmss [((Binding.name "inducts", inducts), [])] |> snd
+ end;
+
+
+
+(**** simplification procedure for showing distinctness of constructors ****)
+
+fun stripT (i, Type ("fun", [_, T])) = stripT (i + 1, T)
+ | stripT p = p;
+
+fun stripC (i, f $ x) = stripC (i + 1, f)
+ | stripC p = p;
+
+val distinctN = "constr_distinct";
+
+fun distinct_rule thy ss tname eq_t = case #distinct (the_datatype thy tname) of
+ FewConstrs thms => Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
+ (EVERY [rtac eq_reflection 1, rtac iffI 1, rtac notE 1,
+ atac 2, resolve_tac thms 1, etac FalseE 1]))
+ | ManyConstrs (thm, simpset) =>
+ let
+ val [In0_inject, In1_inject, In0_not_In1, In1_not_In0] =
+ map (PureThy.get_thm (ThyInfo.the_theory "Datatype" thy))
+ ["In0_inject", "In1_inject", "In0_not_In1", "In1_not_In0"];
+ in
+ Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
+ (EVERY [rtac eq_reflection 1, rtac iffI 1, dtac thm 1,
+ full_simp_tac (Simplifier.inherit_context ss simpset) 1,
+ REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
+ eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1,
+ etac FalseE 1]))
+ end;
+
+fun distinct_proc thy ss (t as Const ("op =", _) $ t1 $ t2) =
+ (case (stripC (0, t1), stripC (0, t2)) of
+ ((i, Const (cname1, T1)), (j, Const (cname2, T2))) =>
+ (case (stripT (0, T1), stripT (0, T2)) of
+ ((i', Type (tname1, _)), (j', Type (tname2, _))) =>
+ if tname1 = tname2 andalso not (cname1 = cname2) andalso i = i' andalso j = j' then
+ (case (get_datatype_descr thy) tname1 of
+ SOME (_, (_, constrs)) => let val cnames = map fst constrs
+ in if cname1 mem cnames andalso cname2 mem cnames then
+ SOME (distinct_rule thy ss tname1
+ (Logic.mk_equals (t, Const ("False", HOLogic.boolT))))
+ else NONE
+ end
+ | NONE => NONE)
+ else NONE
+ | _ => NONE)
+ | _ => NONE)
+ | distinct_proc _ _ _ = NONE;
+
+val distinct_simproc =
+ Simplifier.simproc @{theory HOL} distinctN ["s = t"] distinct_proc;
+
+val dist_ss = HOL_ss addsimprocs [distinct_simproc];
+
+val simproc_setup =
+ Simplifier.map_simpset (fn ss => ss addsimprocs [distinct_simproc]);
+
+
+(**** translation rules for case ****)
+
+fun make_case ctxt = DatatypeCase.make_case
+ (datatype_of_constr (ProofContext.theory_of ctxt)) ctxt;
+
+fun strip_case ctxt = DatatypeCase.strip_case
+ (datatype_of_case (ProofContext.theory_of ctxt));
+
+fun add_case_tr' case_names thy =
+ Sign.add_advanced_trfuns ([], [],
+ map (fn case_name =>
+ let val case_name' = Sign.const_syntax_name thy case_name
+ in (case_name', DatatypeCase.case_tr' datatype_of_case case_name')
+ end) case_names, []) thy;
+
+val trfun_setup =
+ Sign.add_advanced_trfuns ([],
+ [("_case_syntax", DatatypeCase.case_tr true datatype_of_constr)],
+ [], []);
+
+
+(* prepare types *)
+
+fun read_typ thy ((Ts, sorts), str) =
+ let
+ val ctxt = ProofContext.init thy
+ |> fold (Variable.declare_typ o TFree) sorts;
+ val T = Syntax.read_typ ctxt str;
+ in (Ts @ [T], Term.add_tfreesT T sorts) end;
+
+fun cert_typ sign ((Ts, sorts), raw_T) =
+ let
+ val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle
+ TYPE (msg, _, _) => error msg;
+ val sorts' = Term.add_tfreesT T sorts;
+ in (Ts @ [T],
+ case duplicates (op =) (map fst sorts') of
+ [] => sorts'
+ | dups => error ("Inconsistent sort constraints for " ^ commas dups))
+ end;
+
+
+(**** make datatype info ****)
+
+fun make_dt_info alt_names descr sorts induct reccomb_names rec_thms
+ (((((((((i, (_, (tname, _, _))), case_name), case_thms),
+ exhaustion_thm), distinct_thm), inject), nchotomy), case_cong), weak_case_cong) =
+ (tname,
+ {index = i,
+ alt_names = alt_names,
+ descr = descr,
+ sorts = sorts,
+ rec_names = reccomb_names,
+ rec_rewrites = rec_thms,
+ case_name = case_name,
+ case_rewrites = case_thms,
+ induction = induct,
+ exhaustion = exhaustion_thm,
+ distinct = distinct_thm,
+ inject = inject,
+ nchotomy = nchotomy,
+ case_cong = case_cong,
+ weak_case_cong = weak_case_cong});
+
+type rules = {distinct : thm list list,
+ inject : thm list list,
+ exhaustion : thm list,
+ rec_thms : thm list,
+ case_thms : thm list list,
+ split_thms : (thm * thm) list,
+ induction : thm,
+ simps : thm list}
+
+structure DatatypeInterpretation = InterpretationFun
+ (type T = datatype_config * string list val eq: T * T -> bool = eq_snd op =);
+fun interpretation f = DatatypeInterpretation.interpretation (uncurry f);
+
+
+(******************* definitional introduction of datatypes *******************)
+
+fun add_datatype_def (config : datatype_config) new_type_names descr sorts types_syntax constr_syntax dt_info
+ case_names_induct case_names_exhausts thy =
+ let
+ val _ = message config ("Proofs for datatype(s) " ^ commas_quote new_type_names);
+
+ val ((inject, distinct, dist_rewrites, simproc_dists, induct), thy2) = thy |>
+ DatatypeRepProofs.representation_proofs config dt_info new_type_names descr sorts
+ types_syntax constr_syntax case_names_induct;
+
+ val (casedist_thms, thy3) = DatatypeAbsProofs.prove_casedist_thms config new_type_names descr
+ sorts induct case_names_exhausts thy2;
+ val ((reccomb_names, rec_thms), thy4) = DatatypeAbsProofs.prove_primrec_thms
+ config new_type_names descr sorts dt_info inject dist_rewrites
+ (Simplifier.theory_context thy3 dist_ss) induct thy3;
+ val ((case_thms, case_names), thy6) = DatatypeAbsProofs.prove_case_thms
+ config new_type_names descr sorts reccomb_names rec_thms thy4;
+ val (split_thms, thy7) = DatatypeAbsProofs.prove_split_thms config new_type_names
+ descr sorts inject dist_rewrites casedist_thms case_thms thy6;
+ val (nchotomys, thy8) = DatatypeAbsProofs.prove_nchotomys config new_type_names
+ descr sorts casedist_thms thy7;
+ val (case_congs, thy9) = DatatypeAbsProofs.prove_case_congs new_type_names
+ descr sorts nchotomys case_thms thy8;
+ val (weak_case_congs, thy10) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
+ descr sorts thy9;
+
+ val dt_infos = map
+ (make_dt_info (SOME new_type_names) (flat descr) sorts induct reccomb_names rec_thms)
+ ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names ~~ case_thms ~~
+ casedist_thms ~~ simproc_dists ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
+
+ val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
+
+ val thy12 =
+ thy10
+ |> add_case_tr' case_names
+ |> Sign.add_path (space_implode "_" new_type_names)
+ |> add_rules simps case_thms rec_thms inject distinct
+ weak_case_congs (Simplifier.attrib (op addcongs))
+ |> put_dt_infos dt_infos
+ |> add_cases_induct dt_infos induct
+ |> Sign.parent_path
+ |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms) |> snd
+ |> DatatypeInterpretation.data (config, map fst dt_infos);
+ in
+ ({distinct = distinct,
+ inject = inject,
+ exhaustion = casedist_thms,
+ rec_thms = rec_thms,
+ case_thms = case_thms,
+ split_thms = split_thms,
+ induction = induct,
+ simps = simps}, thy12)
+ end;
+
+
+(*********************** declare existing type as datatype *********************)
+
+fun prove_rep_datatype (config : datatype_config) alt_names new_type_names descr sorts induct inject half_distinct thy =
+ let
+ val ((_, [induct']), _) =
+ Variable.importT_thms [induct] (Variable.thm_context induct);
+
+ fun err t = error ("Ill-formed predicate in induction rule: " ^
+ Syntax.string_of_term_global thy t);
+
+ fun get_typ (t as _ $ Var (_, Type (tname, Ts))) =
+ ((tname, map (fst o dest_TFree) Ts) handle TERM _ => err t)
+ | get_typ t = err t;
+ val dtnames = map get_typ (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct')));
+
+ val dt_info = get_datatypes thy;
+
+ val distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct;
+ val (case_names_induct, case_names_exhausts) =
+ (mk_case_names_induct descr, mk_case_names_exhausts descr (map #1 dtnames));
+
+ val _ = message config ("Proofs for datatype(s) " ^ commas_quote new_type_names);
+
+ val (casedist_thms, thy2) = thy |>
+ DatatypeAbsProofs.prove_casedist_thms config new_type_names [descr] sorts induct
+ case_names_exhausts;
+ val ((reccomb_names, rec_thms), thy3) = DatatypeAbsProofs.prove_primrec_thms
+ config new_type_names [descr] sorts dt_info inject distinct
+ (Simplifier.theory_context thy2 dist_ss) induct thy2;
+ val ((case_thms, case_names), thy4) = DatatypeAbsProofs.prove_case_thms
+ config new_type_names [descr] sorts reccomb_names rec_thms thy3;
+ val (split_thms, thy5) = DatatypeAbsProofs.prove_split_thms
+ config new_type_names [descr] sorts inject distinct casedist_thms case_thms thy4;
+ val (nchotomys, thy6) = DatatypeAbsProofs.prove_nchotomys config new_type_names
+ [descr] sorts casedist_thms thy5;
+ val (case_congs, thy7) = DatatypeAbsProofs.prove_case_congs new_type_names
+ [descr] sorts nchotomys case_thms thy6;
+ val (weak_case_congs, thy8) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
+ [descr] sorts thy7;
+
+ val ((_, [induct']), thy10) =
+ thy8
+ |> store_thmss "inject" new_type_names inject
+ ||>> store_thmss "distinct" new_type_names distinct
+ ||> Sign.add_path (space_implode "_" new_type_names)
+ ||>> PureThy.add_thms [((Binding.name "induct", induct), [case_names_induct])];
+
+ val dt_infos = map (make_dt_info alt_names descr sorts induct' reccomb_names rec_thms)
+ ((0 upto length descr - 1) ~~ descr ~~ case_names ~~ case_thms ~~ casedist_thms ~~
+ map FewConstrs distinct ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
+
+ val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
+
+ val thy11 =
+ thy10
+ |> add_case_tr' case_names
+ |> add_rules simps case_thms rec_thms inject distinct
+ weak_case_congs (Simplifier.attrib (op addcongs))
+ |> put_dt_infos dt_infos
+ |> add_cases_induct dt_infos induct'
+ |> Sign.parent_path
+ |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms)
+ |> snd
+ |> DatatypeInterpretation.data (config, map fst dt_infos);
+ in
+ ({distinct = distinct,
+ inject = inject,
+ exhaustion = casedist_thms,
+ rec_thms = rec_thms,
+ case_thms = case_thms,
+ split_thms = split_thms,
+ induction = induct',
+ simps = simps}, thy11)
+ end;
+
+fun gen_rep_datatype prep_term (config : datatype_config) after_qed alt_names raw_ts thy =
+ let
+ fun constr_of_term (Const (c, T)) = (c, T)
+ | constr_of_term t =
+ error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
+ fun no_constr (c, T) = error ("Bad constructor: "
+ ^ Sign.extern_const thy c ^ "::"
+ ^ Syntax.string_of_typ_global thy T);
+ fun type_of_constr (cT as (_, T)) =
+ let
+ val frees = OldTerm.typ_tfrees T;
+ val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) T
+ handle TYPE _ => no_constr cT
+ val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else ();
+ val _ = if length frees <> length vs then no_constr cT else ();
+ in (tyco, (vs, cT)) end;
+
+ val raw_cs = AList.group (op =) (map (type_of_constr o constr_of_term o prep_term thy) raw_ts);
+ val _ = case map_filter (fn (tyco, _) =>
+ if Symtab.defined (get_datatypes thy) tyco then SOME tyco else NONE) raw_cs
+ of [] => ()
+ | tycos => error ("Type(s) " ^ commas (map quote tycos)
+ ^ " already represented inductivly");
+ val raw_vss = maps (map (map snd o fst) o snd) raw_cs;
+ val ms = case distinct (op =) (map length raw_vss)
+ of [n] => 0 upto n - 1
+ | _ => error ("Different types in given constructors");
+ fun inter_sort m = map (fn xs => nth xs m) raw_vss
+ |> Library.foldr1 (Sorts.inter_sort (Sign.classes_of thy))
+ val sorts = map inter_sort ms;
+ val vs = Name.names Name.context Name.aT sorts;
+
+ fun norm_constr (raw_vs, (c, T)) = (c, map_atyps
+ (TFree o (the o AList.lookup (op =) (map fst raw_vs ~~ vs)) o fst o dest_TFree) T);
+
+ val cs = map (apsnd (map norm_constr)) raw_cs;
+ val dtyps_of_typ = map (dtyp_of_typ (map (rpair (map fst vs) o fst) cs))
+ o fst o strip_type;
+ val new_type_names = map Long_Name.base_name (the_default (map fst cs) alt_names);
+
+ fun mk_spec (i, (tyco, constr)) = (i, (tyco,
+ map (DtTFree o fst) vs,
+ (map o apsnd) dtyps_of_typ constr))
+ val descr = map_index mk_spec cs;
+ val injs = DatatypeProp.make_injs [descr] vs;
+ val half_distincts = map snd (DatatypeProp.make_distincts [descr] vs);
+ val ind = DatatypeProp.make_ind [descr] vs;
+ val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts];
+
+ fun after_qed' raw_thms =
+ let
+ val [[[induct]], injs, half_distincts] =
+ unflat rules (map Drule.zero_var_indexes_list raw_thms);
+ (*FIXME somehow dubious*)
+ in
+ ProofContext.theory_result
+ (prove_rep_datatype config alt_names new_type_names descr vs induct injs half_distincts)
+ #-> after_qed
+ end;
+ in
+ thy
+ |> ProofContext.init
+ |> Proof.theorem_i NONE after_qed' ((map o map) (rpair []) (flat rules))
+ end;
+
+val rep_datatype = gen_rep_datatype Sign.cert_term;
+val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global default_datatype_config (K I);
+
+
+
+(******************************** add datatype ********************************)
+
+fun gen_add_datatype prep_typ (config : datatype_config) new_type_names dts thy =
+ let
+ val _ = Theory.requires thy "Datatype" "datatype definitions";
+
+ (* this theory is used just for parsing *)
+
+ val tmp_thy = thy |>
+ Theory.copy |>
+ Sign.add_types (map (fn (tvs, tname, mx, _) =>
+ (tname, length tvs, mx)) dts);
+
+ val (tyvars, _, _, _)::_ = dts;
+ val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
+ let val full_tname = Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname)
+ in (case duplicates (op =) tvs of
+ [] => if eq_set (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
+ else error ("Mutually recursive datatypes must have same type parameters")
+ | dups => error ("Duplicate parameter(s) for datatype " ^ quote (Binding.str_of tname) ^
+ " : " ^ commas dups))
+ end) dts);
+
+ val _ = (case duplicates (op =) (map fst new_dts) @ duplicates (op =) new_type_names of
+ [] => () | dups => error ("Duplicate datatypes: " ^ commas dups));
+
+ fun prep_dt_spec ((tvs, tname, mx, constrs), tname') (dts', constr_syntax, sorts, i) =
+ let
+ fun prep_constr (cname, cargs, mx') (constrs, constr_syntax', sorts') =
+ let
+ val (cargs', sorts'') = Library.foldl (prep_typ tmp_thy) (([], sorts'), cargs);
+ val _ = (case fold (curry OldTerm.add_typ_tfree_names) cargs' [] \\ tvs of
+ [] => ()
+ | vs => error ("Extra type variables on rhs: " ^ commas vs))
+ in (constrs @ [((if #flat_names config then Sign.full_name tmp_thy else
+ Sign.full_name_path tmp_thy tname')
+ (Binding.map_name (Syntax.const_name mx') cname),
+ map (dtyp_of_typ new_dts) cargs')],
+ constr_syntax' @ [(cname, mx')], sorts'')
+ end handle ERROR msg => cat_error msg
+ ("The error above occured in constructor " ^ quote (Binding.str_of cname) ^
+ " of datatype " ^ quote (Binding.str_of tname));
+
+ val (constrs', constr_syntax', sorts') =
+ fold prep_constr constrs ([], [], sorts)
+
+ in
+ case duplicates (op =) (map fst constrs') of
+ [] =>
+ (dts' @ [(i, (Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname),
+ map DtTFree tvs, constrs'))],
+ constr_syntax @ [constr_syntax'], sorts', i + 1)
+ | dups => error ("Duplicate constructors " ^ commas dups ^
+ " in datatype " ^ quote (Binding.str_of tname))
+ end;
+
+ val (dts', constr_syntax, sorts', i) =
+ fold prep_dt_spec (dts ~~ new_type_names) ([], [], [], 0);
+ val sorts = sorts' @ (map (rpair (Sign.defaultS tmp_thy)) (tyvars \\ map fst sorts'));
+ val dt_info = get_datatypes thy;
+ val (descr, _) = unfold_datatypes tmp_thy dts' sorts dt_info dts' i;
+ val _ = check_nonempty descr handle (exn as Datatype_Empty s) =>
+ if #strict config then error ("Nonemptiness check failed for datatype " ^ s)
+ else raise exn;
+
+ val descr' = flat descr;
+ val case_names_induct = mk_case_names_induct descr';
+ val case_names_exhausts = mk_case_names_exhausts descr' (map #1 new_dts);
+ in
+ add_datatype_def
+ (config : datatype_config) new_type_names descr sorts types_syntax constr_syntax dt_info
+ case_names_induct case_names_exhausts thy
+ end;
+
+val add_datatype = gen_add_datatype cert_typ;
+val add_datatype_cmd = gen_add_datatype read_typ default_datatype_config;
+
+
+
+(** package setup **)
+
+(* setup theory *)
+
+val setup =
+ DatatypeRepProofs.distinctness_limit_setup #>
+ simproc_setup #>
+ trfun_setup #>
+ DatatypeInterpretation.init;
+
+
+(* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val datatype_decl =
+ Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.binding -- P.opt_infix --
+ (P.$$$ "=" |-- P.enum1 "|" (P.binding -- Scan.repeat P.typ -- P.opt_mixfix));
+
+fun mk_datatype args =
+ let
+ val names = map
+ (fn ((((NONE, _), t), _), _) => Binding.name_of t | ((((SOME t, _), _), _), _) => t) args;
+ val specs = map (fn ((((_, vs), t), mx), cons) =>
+ (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
+ in snd o add_datatype_cmd names specs end;
+
+val _ =
+ OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl
+ (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
+
+val _ =
+ OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_goal
+ (Scan.option (P.$$$ "(" |-- Scan.repeat1 P.name --| P.$$$ ")") -- Scan.repeat1 P.term
+ >> (fn (alt_names, ts) => Toplevel.print
+ o Toplevel.theory_to_proof (rep_datatype_cmd alt_names ts)));
+
+end;
+
+
+(* document antiquotation *)
+
+val _ = ThyOutput.antiquotation "datatype" Args.tyname
+ (fn {source = src, context = ctxt, ...} => fn dtco =>
+ let
+ val thy = ProofContext.theory_of ctxt;
+ val (vs, cos) = the_datatype_spec thy dtco;
+ val ty = Type (dtco, map TFree vs);
+ fun pretty_typ_bracket (ty as Type (_, _ :: _)) =
+ Pretty.enclose "(" ")" [Syntax.pretty_typ ctxt ty]
+ | pretty_typ_bracket ty =
+ Syntax.pretty_typ ctxt ty;
+ fun pretty_constr (co, tys) =
+ (Pretty.block o Pretty.breaks)
+ (Syntax.pretty_term ctxt (Const (co, tys ---> ty)) ::
+ map pretty_typ_bracket tys);
+ val pretty_datatype =
+ Pretty.block
+ (Pretty.command "datatype" :: Pretty.brk 1 ::
+ Syntax.pretty_typ ctxt ty ::
+ Pretty.str " =" :: Pretty.brk 1 ::
+ flat (separate [Pretty.brk 1, Pretty.str "| "]
+ (map (single o pretty_constr) cos)));
+ in ThyOutput.output (ThyOutput.maybe_pretty_source (K pretty_datatype) src [()]) end);
+
+end;
+
--- a/src/HOL/Tools/datatype_package/datatype_abs_proofs.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Tools/datatype_package/datatype_abs_proofs.ML Fri Jun 19 17:23:21 2009 +0200
@@ -155,7 +155,7 @@
(([], 0), descr' ~~ recTs ~~ rec_sets');
val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
- InductivePackage.add_inductive_global (serial_string ())
+ Inductive.add_inductive_global (serial_string ())
{quiet_mode = #quiet config, verbose = false, kind = Thm.internalK,
alt_name = Binding.name big_rec_name', coind = false, no_elim = false, no_ind = true,
skip_mono = true, fork_mono = false}
--- a/src/HOL/Tools/datatype_package/datatype_codegen.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Tools/datatype_package/datatype_codegen.ML Fri Jun 19 17:23:21 2009 +0200
@@ -276,12 +276,12 @@
fun datatype_codegen thy defs dep module brack t gr = (case strip_comb t of
(c as Const (s, T), ts) =>
- (case DatatypePackage.datatype_of_case thy s of
+ (case Datatype.datatype_of_case thy s of
SOME {index, descr, ...} =>
if is_some (get_assoc_code thy (s, T)) then NONE else
SOME (pretty_case thy defs dep module brack
(#3 (the (AList.lookup op = descr index))) c ts gr )
- | NONE => case (DatatypePackage.datatype_of_constr thy s, strip_type T) of
+ | NONE => case (Datatype.datatype_of_constr thy s, strip_type T) of
(SOME {index, descr, ...}, (_, U as Type (tyname, _))) =>
if is_some (get_assoc_code thy (s, T)) then NONE else
let
@@ -296,7 +296,7 @@
| _ => NONE);
fun datatype_tycodegen thy defs dep module brack (Type (s, Ts)) gr =
- (case DatatypePackage.get_datatype thy s of
+ (case Datatype.get_datatype thy s of
NONE => NONE
| SOME {descr, sorts, ...} =>
if is_some (get_assoc_type thy s) then NONE else
@@ -331,7 +331,7 @@
fun mk_case_cert thy tyco =
let
val raw_thms =
- (#case_rewrites o DatatypePackage.the_datatype thy) tyco;
+ (#case_rewrites o Datatype.the_datatype thy) tyco;
val thms as hd_thm :: _ = raw_thms
|> Conjunction.intr_balanced
|> Thm.unvarify
@@ -364,8 +364,8 @@
fun mk_eq_eqns thy dtco =
let
- val (vs, cos) = DatatypePackage.the_datatype_spec thy dtco;
- val { descr, index, inject = inject_thms, ... } = DatatypePackage.the_datatype thy dtco;
+ val (vs, cos) = Datatype.the_datatype_spec thy dtco;
+ val { descr, index, inject = inject_thms, ... } = Datatype.the_datatype thy dtco;
val ty = Type (dtco, map TFree vs);
fun mk_eq (t1, t2) = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
$ t1 $ t2;
@@ -383,7 +383,7 @@
val refl = HOLogic.mk_Trueprop (true_eq (Free ("x", ty), Free ("x", ty)));
val simpset = Simplifier.context (ProofContext.init thy) (HOL_basic_ss
addsimps (map Simpdata.mk_eq (@{thm eq} :: @{thm eq_True} :: inject_thms))
- addsimprocs [DatatypePackage.distinct_simproc]);
+ addsimprocs [Datatype.distinct_simproc]);
fun prove prop = SkipProof.prove_global thy [] [] prop (K (ALLGOALS (simp_tac simpset)))
|> Simpdata.mk_eq;
in map (rpair true o prove) (triv_injects @ injects @ distincts) @ [(prove refl, false)] end;
@@ -428,11 +428,11 @@
fun add_all_code config dtcos thy =
let
- val (vs :: _, coss) = (split_list o map (DatatypePackage.the_datatype_spec thy)) dtcos;
+ val (vs :: _, coss) = (split_list o map (Datatype.the_datatype_spec thy)) dtcos;
val any_css = map2 (mk_constr_consts thy vs) dtcos coss;
val css = if exists is_none any_css then []
else map_filter I any_css;
- val case_rewrites = maps (#case_rewrites o DatatypePackage.the_datatype thy) dtcos;
+ val case_rewrites = maps (#case_rewrites o Datatype.the_datatype thy) dtcos;
val certs = map (mk_case_cert thy) dtcos;
in
if null css then thy
@@ -450,6 +450,6 @@
val setup =
add_codegen "datatype" datatype_codegen
#> add_tycodegen "datatype" datatype_tycodegen
- #> DatatypePackage.interpretation add_all_code
+ #> Datatype.interpretation add_all_code
end;
--- a/src/HOL/Tools/datatype_package/datatype_package.ML Thu Jun 18 18:31:14 2009 -0700
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,705 +0,0 @@
-(* Title: HOL/Tools/datatype_package.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Datatype package for Isabelle/HOL.
-*)
-
-signature DATATYPE_PACKAGE =
-sig
- type datatype_config = DatatypeAux.datatype_config
- type datatype_info = DatatypeAux.datatype_info
- type descr = DatatypeAux.descr
- val get_datatypes : theory -> datatype_info Symtab.table
- val get_datatype : theory -> string -> datatype_info option
- val the_datatype : theory -> string -> datatype_info
- val datatype_of_constr : theory -> string -> datatype_info option
- val datatype_of_case : theory -> string -> datatype_info option
- val the_datatype_spec : theory -> string -> (string * sort) list * (string * typ list) list
- val the_datatype_descr : theory -> string list
- -> descr * (string * sort) list * string list
- * (string list * string list) * (typ list * typ list)
- val get_datatype_constrs : theory -> string -> (string * typ) list option
- val distinct_simproc : simproc
- val make_case : Proof.context -> bool -> string list -> term ->
- (term * term) list -> term * (term * (int * bool)) list
- val strip_case : Proof.context -> bool -> term -> (term * (term * term) list) option
- val read_typ: theory ->
- (typ list * (string * sort) list) * string -> typ list * (string * sort) list
- val interpretation : (datatype_config -> string list -> theory -> theory) -> theory -> theory
- type rules = {distinct : thm list list,
- inject : thm list list,
- exhaustion : thm list,
- rec_thms : thm list,
- case_thms : thm list list,
- split_thms : (thm * thm) list,
- induction : thm,
- simps : thm list}
- val rep_datatype : datatype_config -> (rules -> Proof.context -> Proof.context)
- -> string list option -> term list -> theory -> Proof.state;
- val rep_datatype_cmd : string list option -> string list -> theory -> Proof.state
- val add_datatype : datatype_config -> string list -> (string list * binding * mixfix *
- (binding * typ list * mixfix) list) list -> theory -> rules * theory
- val add_datatype_cmd : string list -> (string list * binding * mixfix *
- (binding * string list * mixfix) list) list -> theory -> rules * theory
- val setup: theory -> theory
- val print_datatypes : theory -> unit
-end;
-
-structure DatatypePackage : DATATYPE_PACKAGE =
-struct
-
-open DatatypeAux;
-
-
-(* theory data *)
-
-structure DatatypesData = TheoryDataFun
-(
- type T =
- {types: datatype_info Symtab.table,
- constrs: datatype_info Symtab.table,
- cases: datatype_info Symtab.table};
-
- val empty =
- {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty};
- val copy = I;
- val extend = I;
- fun merge _
- ({types = types1, constrs = constrs1, cases = cases1},
- {types = types2, constrs = constrs2, cases = cases2}) =
- {types = Symtab.merge (K true) (types1, types2),
- constrs = Symtab.merge (K true) (constrs1, constrs2),
- cases = Symtab.merge (K true) (cases1, cases2)};
-);
-
-val get_datatypes = #types o DatatypesData.get;
-val map_datatypes = DatatypesData.map;
-
-fun print_datatypes thy =
- Pretty.writeln (Pretty.strs ("datatypes:" ::
- map #1 (NameSpace.extern_table (Sign.type_space thy, get_datatypes thy))));
-
-
-(** theory information about datatypes **)
-
-fun put_dt_infos (dt_infos : (string * datatype_info) list) =
- map_datatypes (fn {types, constrs, cases} =>
- {types = fold Symtab.update dt_infos types,
- constrs = fold Symtab.default (*conservative wrt. overloaded constructors*)
- (maps (fn (_, info as {descr, index, ...}) => map (rpair info o fst)
- (#3 (the (AList.lookup op = descr index)))) dt_infos) constrs,
- cases = fold Symtab.update
- (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)
- cases});
-
-val get_datatype = Symtab.lookup o get_datatypes;
-
-fun the_datatype thy name = (case get_datatype thy name of
- SOME info => info
- | NONE => error ("Unknown datatype " ^ quote name));
-
-val datatype_of_constr = Symtab.lookup o #constrs o DatatypesData.get;
-val datatype_of_case = Symtab.lookup o #cases o DatatypesData.get;
-
-fun get_datatype_descr thy dtco =
- get_datatype thy dtco
- |> Option.map (fn info as { descr, index, ... } =>
- (info, (((fn SOME (_, dtys, cos) => (dtys, cos)) o AList.lookup (op =) descr) index)));
-
-fun the_datatype_spec thy dtco =
- let
- val info as { descr, index, sorts = raw_sorts, ... } = the_datatype thy dtco;
- val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr index;
- val sorts = map ((fn v => (v, (the o AList.lookup (op =) raw_sorts) v))
- o DatatypeAux.dest_DtTFree) dtys;
- val cos = map
- (fn (co, tys) => (co, map (DatatypeAux.typ_of_dtyp descr sorts) tys)) raw_cos;
- in (sorts, cos) end;
-
-fun the_datatype_descr thy (raw_tycos as raw_tyco :: _) =
- let
- val info = the_datatype thy raw_tyco;
- val descr = #descr info;
-
- val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr (#index info);
- val vs = map ((fn v => (v, (the o AList.lookup (op =) (#sorts info)) v))
- o dest_DtTFree) dtys;
-
- fun is_DtTFree (DtTFree _) = true
- | is_DtTFree _ = false
- val k = find_index (fn (_, (_, dTs, _)) => not (forall is_DtTFree dTs)) descr;
- val protoTs as (dataTs, _) = chop k descr
- |> (pairself o map) (fn (_, (tyco, dTs, _)) => (tyco, map (typ_of_dtyp descr vs) dTs));
-
- val tycos = map fst dataTs;
- val _ = if gen_eq_set (op =) (tycos, raw_tycos) then ()
- else error ("Type constructors " ^ commas (map quote raw_tycos)
- ^ "do not belong exhaustively to one mutual recursive datatype");
-
- val (Ts, Us) = (pairself o map) Type protoTs;
-
- val names = map Long_Name.base_name (the_default tycos (#alt_names info));
- val (auxnames, _) = Name.make_context names
- |> fold_map (yield_singleton Name.variants o name_of_typ) Us
-
- in (descr, vs, tycos, (names, auxnames), (Ts, Us)) end;
-
-fun get_datatype_constrs thy dtco =
- case try (the_datatype_spec thy) dtco
- of SOME (sorts, cos) =>
- let
- fun subst (v, sort) = TVar ((v, 0), sort);
- fun subst_ty (TFree v) = subst v
- | subst_ty ty = ty;
- val dty = Type (dtco, map subst sorts);
- fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty);
- in SOME (map mk_co cos) end
- | NONE => NONE;
-
-
-(** induct method setup **)
-
-(* case names *)
-
-local
-
-fun dt_recs (DtTFree _) = []
- | dt_recs (DtType (_, dts)) = maps dt_recs dts
- | dt_recs (DtRec i) = [i];
-
-fun dt_cases (descr: descr) (_, args, constrs) =
- let
- fun the_bname i = Long_Name.base_name (#1 (the (AList.lookup (op =) descr i)));
- val bnames = map the_bname (distinct (op =) (maps dt_recs args));
- in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end;
-
-
-fun induct_cases descr =
- DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr));
-
-fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i));
-
-in
-
-fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
-
-fun mk_case_names_exhausts descr new =
- map (RuleCases.case_names o exhaust_cases descr o #1)
- (filter (fn ((_, (name, _, _))) => member (op =) new name) descr);
-
-end;
-
-fun add_rules simps case_thms rec_thms inject distinct
- weak_case_congs cong_att =
- PureThy.add_thmss [((Binding.name "simps", simps), []),
- ((Binding.empty, flat case_thms @
- flat distinct @ rec_thms), [Simplifier.simp_add]),
- ((Binding.empty, rec_thms), [Code.add_default_eqn_attribute]),
- ((Binding.empty, flat inject), [iff_add]),
- ((Binding.empty, map (fn th => th RS notE) (flat distinct)), [Classical.safe_elim NONE]),
- ((Binding.empty, weak_case_congs), [cong_att])]
- #> snd;
-
-
-(* add_cases_induct *)
-
-fun add_cases_induct infos induction thy =
- let
- val inducts = ProjectRule.projections (ProofContext.init thy) induction;
-
- fun named_rules (name, {index, exhaustion, ...}: datatype_info) =
- [((Binding.empty, nth inducts index), [Induct.induct_type name]),
- ((Binding.empty, exhaustion), [Induct.cases_type name])];
- fun unnamed_rule i =
- ((Binding.empty, nth inducts i), [Thm.kind_internal, Induct.induct_type ""]);
- in
- thy |> PureThy.add_thms
- (maps named_rules infos @
- map unnamed_rule (length infos upto length inducts - 1)) |> snd
- |> PureThy.add_thmss [((Binding.name "inducts", inducts), [])] |> snd
- end;
-
-
-
-(**** simplification procedure for showing distinctness of constructors ****)
-
-fun stripT (i, Type ("fun", [_, T])) = stripT (i + 1, T)
- | stripT p = p;
-
-fun stripC (i, f $ x) = stripC (i + 1, f)
- | stripC p = p;
-
-val distinctN = "constr_distinct";
-
-fun distinct_rule thy ss tname eq_t = case #distinct (the_datatype thy tname) of
- FewConstrs thms => Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
- (EVERY [rtac eq_reflection 1, rtac iffI 1, rtac notE 1,
- atac 2, resolve_tac thms 1, etac FalseE 1]))
- | ManyConstrs (thm, simpset) =>
- let
- val [In0_inject, In1_inject, In0_not_In1, In1_not_In0] =
- map (PureThy.get_thm (ThyInfo.the_theory "Datatype" thy))
- ["In0_inject", "In1_inject", "In0_not_In1", "In1_not_In0"];
- in
- Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
- (EVERY [rtac eq_reflection 1, rtac iffI 1, dtac thm 1,
- full_simp_tac (Simplifier.inherit_context ss simpset) 1,
- REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
- eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1,
- etac FalseE 1]))
- end;
-
-fun distinct_proc thy ss (t as Const ("op =", _) $ t1 $ t2) =
- (case (stripC (0, t1), stripC (0, t2)) of
- ((i, Const (cname1, T1)), (j, Const (cname2, T2))) =>
- (case (stripT (0, T1), stripT (0, T2)) of
- ((i', Type (tname1, _)), (j', Type (tname2, _))) =>
- if tname1 = tname2 andalso not (cname1 = cname2) andalso i = i' andalso j = j' then
- (case (get_datatype_descr thy) tname1 of
- SOME (_, (_, constrs)) => let val cnames = map fst constrs
- in if cname1 mem cnames andalso cname2 mem cnames then
- SOME (distinct_rule thy ss tname1
- (Logic.mk_equals (t, Const ("False", HOLogic.boolT))))
- else NONE
- end
- | NONE => NONE)
- else NONE
- | _ => NONE)
- | _ => NONE)
- | distinct_proc _ _ _ = NONE;
-
-val distinct_simproc =
- Simplifier.simproc @{theory HOL} distinctN ["s = t"] distinct_proc;
-
-val dist_ss = HOL_ss addsimprocs [distinct_simproc];
-
-val simproc_setup =
- Simplifier.map_simpset (fn ss => ss addsimprocs [distinct_simproc]);
-
-
-(**** translation rules for case ****)
-
-fun make_case ctxt = DatatypeCase.make_case
- (datatype_of_constr (ProofContext.theory_of ctxt)) ctxt;
-
-fun strip_case ctxt = DatatypeCase.strip_case
- (datatype_of_case (ProofContext.theory_of ctxt));
-
-fun add_case_tr' case_names thy =
- Sign.add_advanced_trfuns ([], [],
- map (fn case_name =>
- let val case_name' = Sign.const_syntax_name thy case_name
- in (case_name', DatatypeCase.case_tr' datatype_of_case case_name')
- end) case_names, []) thy;
-
-val trfun_setup =
- Sign.add_advanced_trfuns ([],
- [("_case_syntax", DatatypeCase.case_tr true datatype_of_constr)],
- [], []);
-
-
-(* prepare types *)
-
-fun read_typ thy ((Ts, sorts), str) =
- let
- val ctxt = ProofContext.init thy
- |> fold (Variable.declare_typ o TFree) sorts;
- val T = Syntax.read_typ ctxt str;
- in (Ts @ [T], Term.add_tfreesT T sorts) end;
-
-fun cert_typ sign ((Ts, sorts), raw_T) =
- let
- val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle
- TYPE (msg, _, _) => error msg;
- val sorts' = Term.add_tfreesT T sorts;
- in (Ts @ [T],
- case duplicates (op =) (map fst sorts') of
- [] => sorts'
- | dups => error ("Inconsistent sort constraints for " ^ commas dups))
- end;
-
-
-(**** make datatype info ****)
-
-fun make_dt_info alt_names descr sorts induct reccomb_names rec_thms
- (((((((((i, (_, (tname, _, _))), case_name), case_thms),
- exhaustion_thm), distinct_thm), inject), nchotomy), case_cong), weak_case_cong) =
- (tname,
- {index = i,
- alt_names = alt_names,
- descr = descr,
- sorts = sorts,
- rec_names = reccomb_names,
- rec_rewrites = rec_thms,
- case_name = case_name,
- case_rewrites = case_thms,
- induction = induct,
- exhaustion = exhaustion_thm,
- distinct = distinct_thm,
- inject = inject,
- nchotomy = nchotomy,
- case_cong = case_cong,
- weak_case_cong = weak_case_cong});
-
-type rules = {distinct : thm list list,
- inject : thm list list,
- exhaustion : thm list,
- rec_thms : thm list,
- case_thms : thm list list,
- split_thms : (thm * thm) list,
- induction : thm,
- simps : thm list}
-
-structure DatatypeInterpretation = InterpretationFun
- (type T = datatype_config * string list val eq: T * T -> bool = eq_snd op =);
-fun interpretation f = DatatypeInterpretation.interpretation (uncurry f);
-
-
-(******************* definitional introduction of datatypes *******************)
-
-fun add_datatype_def (config : datatype_config) new_type_names descr sorts types_syntax constr_syntax dt_info
- case_names_induct case_names_exhausts thy =
- let
- val _ = message config ("Proofs for datatype(s) " ^ commas_quote new_type_names);
-
- val ((inject, distinct, dist_rewrites, simproc_dists, induct), thy2) = thy |>
- DatatypeRepProofs.representation_proofs config dt_info new_type_names descr sorts
- types_syntax constr_syntax case_names_induct;
-
- val (casedist_thms, thy3) = DatatypeAbsProofs.prove_casedist_thms config new_type_names descr
- sorts induct case_names_exhausts thy2;
- val ((reccomb_names, rec_thms), thy4) = DatatypeAbsProofs.prove_primrec_thms
- config new_type_names descr sorts dt_info inject dist_rewrites
- (Simplifier.theory_context thy3 dist_ss) induct thy3;
- val ((case_thms, case_names), thy6) = DatatypeAbsProofs.prove_case_thms
- config new_type_names descr sorts reccomb_names rec_thms thy4;
- val (split_thms, thy7) = DatatypeAbsProofs.prove_split_thms config new_type_names
- descr sorts inject dist_rewrites casedist_thms case_thms thy6;
- val (nchotomys, thy8) = DatatypeAbsProofs.prove_nchotomys config new_type_names
- descr sorts casedist_thms thy7;
- val (case_congs, thy9) = DatatypeAbsProofs.prove_case_congs new_type_names
- descr sorts nchotomys case_thms thy8;
- val (weak_case_congs, thy10) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
- descr sorts thy9;
-
- val dt_infos = map
- (make_dt_info (SOME new_type_names) (flat descr) sorts induct reccomb_names rec_thms)
- ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names ~~ case_thms ~~
- casedist_thms ~~ simproc_dists ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
-
- val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
-
- val thy12 =
- thy10
- |> add_case_tr' case_names
- |> Sign.add_path (space_implode "_" new_type_names)
- |> add_rules simps case_thms rec_thms inject distinct
- weak_case_congs (Simplifier.attrib (op addcongs))
- |> put_dt_infos dt_infos
- |> add_cases_induct dt_infos induct
- |> Sign.parent_path
- |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms) |> snd
- |> DatatypeInterpretation.data (config, map fst dt_infos);
- in
- ({distinct = distinct,
- inject = inject,
- exhaustion = casedist_thms,
- rec_thms = rec_thms,
- case_thms = case_thms,
- split_thms = split_thms,
- induction = induct,
- simps = simps}, thy12)
- end;
-
-
-(*********************** declare existing type as datatype *********************)
-
-fun prove_rep_datatype (config : datatype_config) alt_names new_type_names descr sorts induct inject half_distinct thy =
- let
- val ((_, [induct']), _) =
- Variable.importT_thms [induct] (Variable.thm_context induct);
-
- fun err t = error ("Ill-formed predicate in induction rule: " ^
- Syntax.string_of_term_global thy t);
-
- fun get_typ (t as _ $ Var (_, Type (tname, Ts))) =
- ((tname, map (fst o dest_TFree) Ts) handle TERM _ => err t)
- | get_typ t = err t;
- val dtnames = map get_typ (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct')));
-
- val dt_info = get_datatypes thy;
-
- val distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct;
- val (case_names_induct, case_names_exhausts) =
- (mk_case_names_induct descr, mk_case_names_exhausts descr (map #1 dtnames));
-
- val _ = message config ("Proofs for datatype(s) " ^ commas_quote new_type_names);
-
- val (casedist_thms, thy2) = thy |>
- DatatypeAbsProofs.prove_casedist_thms config new_type_names [descr] sorts induct
- case_names_exhausts;
- val ((reccomb_names, rec_thms), thy3) = DatatypeAbsProofs.prove_primrec_thms
- config new_type_names [descr] sorts dt_info inject distinct
- (Simplifier.theory_context thy2 dist_ss) induct thy2;
- val ((case_thms, case_names), thy4) = DatatypeAbsProofs.prove_case_thms
- config new_type_names [descr] sorts reccomb_names rec_thms thy3;
- val (split_thms, thy5) = DatatypeAbsProofs.prove_split_thms
- config new_type_names [descr] sorts inject distinct casedist_thms case_thms thy4;
- val (nchotomys, thy6) = DatatypeAbsProofs.prove_nchotomys config new_type_names
- [descr] sorts casedist_thms thy5;
- val (case_congs, thy7) = DatatypeAbsProofs.prove_case_congs new_type_names
- [descr] sorts nchotomys case_thms thy6;
- val (weak_case_congs, thy8) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
- [descr] sorts thy7;
-
- val ((_, [induct']), thy10) =
- thy8
- |> store_thmss "inject" new_type_names inject
- ||>> store_thmss "distinct" new_type_names distinct
- ||> Sign.add_path (space_implode "_" new_type_names)
- ||>> PureThy.add_thms [((Binding.name "induct", induct), [case_names_induct])];
-
- val dt_infos = map (make_dt_info alt_names descr sorts induct' reccomb_names rec_thms)
- ((0 upto length descr - 1) ~~ descr ~~ case_names ~~ case_thms ~~ casedist_thms ~~
- map FewConstrs distinct ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
-
- val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
-
- val thy11 =
- thy10
- |> add_case_tr' case_names
- |> add_rules simps case_thms rec_thms inject distinct
- weak_case_congs (Simplifier.attrib (op addcongs))
- |> put_dt_infos dt_infos
- |> add_cases_induct dt_infos induct'
- |> Sign.parent_path
- |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms)
- |> snd
- |> DatatypeInterpretation.data (config, map fst dt_infos);
- in
- ({distinct = distinct,
- inject = inject,
- exhaustion = casedist_thms,
- rec_thms = rec_thms,
- case_thms = case_thms,
- split_thms = split_thms,
- induction = induct',
- simps = simps}, thy11)
- end;
-
-fun gen_rep_datatype prep_term (config : datatype_config) after_qed alt_names raw_ts thy =
- let
- fun constr_of_term (Const (c, T)) = (c, T)
- | constr_of_term t =
- error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
- fun no_constr (c, T) = error ("Bad constructor: "
- ^ Sign.extern_const thy c ^ "::"
- ^ Syntax.string_of_typ_global thy T);
- fun type_of_constr (cT as (_, T)) =
- let
- val frees = OldTerm.typ_tfrees T;
- val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) T
- handle TYPE _ => no_constr cT
- val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else ();
- val _ = if length frees <> length vs then no_constr cT else ();
- in (tyco, (vs, cT)) end;
-
- val raw_cs = AList.group (op =) (map (type_of_constr o constr_of_term o prep_term thy) raw_ts);
- val _ = case map_filter (fn (tyco, _) =>
- if Symtab.defined (get_datatypes thy) tyco then SOME tyco else NONE) raw_cs
- of [] => ()
- | tycos => error ("Type(s) " ^ commas (map quote tycos)
- ^ " already represented inductivly");
- val raw_vss = maps (map (map snd o fst) o snd) raw_cs;
- val ms = case distinct (op =) (map length raw_vss)
- of [n] => 0 upto n - 1
- | _ => error ("Different types in given constructors");
- fun inter_sort m = map (fn xs => nth xs m) raw_vss
- |> Library.foldr1 (Sorts.inter_sort (Sign.classes_of thy))
- val sorts = map inter_sort ms;
- val vs = Name.names Name.context Name.aT sorts;
-
- fun norm_constr (raw_vs, (c, T)) = (c, map_atyps
- (TFree o (the o AList.lookup (op =) (map fst raw_vs ~~ vs)) o fst o dest_TFree) T);
-
- val cs = map (apsnd (map norm_constr)) raw_cs;
- val dtyps_of_typ = map (dtyp_of_typ (map (rpair (map fst vs) o fst) cs))
- o fst o strip_type;
- val new_type_names = map Long_Name.base_name (the_default (map fst cs) alt_names);
-
- fun mk_spec (i, (tyco, constr)) = (i, (tyco,
- map (DtTFree o fst) vs,
- (map o apsnd) dtyps_of_typ constr))
- val descr = map_index mk_spec cs;
- val injs = DatatypeProp.make_injs [descr] vs;
- val half_distincts = map snd (DatatypeProp.make_distincts [descr] vs);
- val ind = DatatypeProp.make_ind [descr] vs;
- val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts];
-
- fun after_qed' raw_thms =
- let
- val [[[induct]], injs, half_distincts] =
- unflat rules (map Drule.zero_var_indexes_list raw_thms);
- (*FIXME somehow dubious*)
- in
- ProofContext.theory_result
- (prove_rep_datatype config alt_names new_type_names descr vs induct injs half_distincts)
- #-> after_qed
- end;
- in
- thy
- |> ProofContext.init
- |> Proof.theorem_i NONE after_qed' ((map o map) (rpair []) (flat rules))
- end;
-
-val rep_datatype = gen_rep_datatype Sign.cert_term;
-val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global default_datatype_config (K I);
-
-
-
-(******************************** add datatype ********************************)
-
-fun gen_add_datatype prep_typ (config : datatype_config) new_type_names dts thy =
- let
- val _ = Theory.requires thy "Datatype" "datatype definitions";
-
- (* this theory is used just for parsing *)
-
- val tmp_thy = thy |>
- Theory.copy |>
- Sign.add_types (map (fn (tvs, tname, mx, _) =>
- (tname, length tvs, mx)) dts);
-
- val (tyvars, _, _, _)::_ = dts;
- val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
- let val full_tname = Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname)
- in (case duplicates (op =) tvs of
- [] => if eq_set (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
- else error ("Mutually recursive datatypes must have same type parameters")
- | dups => error ("Duplicate parameter(s) for datatype " ^ quote (Binding.str_of tname) ^
- " : " ^ commas dups))
- end) dts);
-
- val _ = (case duplicates (op =) (map fst new_dts) @ duplicates (op =) new_type_names of
- [] => () | dups => error ("Duplicate datatypes: " ^ commas dups));
-
- fun prep_dt_spec ((tvs, tname, mx, constrs), tname') (dts', constr_syntax, sorts, i) =
- let
- fun prep_constr (cname, cargs, mx') (constrs, constr_syntax', sorts') =
- let
- val (cargs', sorts'') = Library.foldl (prep_typ tmp_thy) (([], sorts'), cargs);
- val _ = (case fold (curry OldTerm.add_typ_tfree_names) cargs' [] \\ tvs of
- [] => ()
- | vs => error ("Extra type variables on rhs: " ^ commas vs))
- in (constrs @ [((if #flat_names config then Sign.full_name tmp_thy else
- Sign.full_name_path tmp_thy tname')
- (Binding.map_name (Syntax.const_name mx') cname),
- map (dtyp_of_typ new_dts) cargs')],
- constr_syntax' @ [(cname, mx')], sorts'')
- end handle ERROR msg => cat_error msg
- ("The error above occured in constructor " ^ quote (Binding.str_of cname) ^
- " of datatype " ^ quote (Binding.str_of tname));
-
- val (constrs', constr_syntax', sorts') =
- fold prep_constr constrs ([], [], sorts)
-
- in
- case duplicates (op =) (map fst constrs') of
- [] =>
- (dts' @ [(i, (Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname),
- map DtTFree tvs, constrs'))],
- constr_syntax @ [constr_syntax'], sorts', i + 1)
- | dups => error ("Duplicate constructors " ^ commas dups ^
- " in datatype " ^ quote (Binding.str_of tname))
- end;
-
- val (dts', constr_syntax, sorts', i) =
- fold prep_dt_spec (dts ~~ new_type_names) ([], [], [], 0);
- val sorts = sorts' @ (map (rpair (Sign.defaultS tmp_thy)) (tyvars \\ map fst sorts'));
- val dt_info = get_datatypes thy;
- val (descr, _) = unfold_datatypes tmp_thy dts' sorts dt_info dts' i;
- val _ = check_nonempty descr handle (exn as Datatype_Empty s) =>
- if #strict config then error ("Nonemptiness check failed for datatype " ^ s)
- else raise exn;
-
- val descr' = flat descr;
- val case_names_induct = mk_case_names_induct descr';
- val case_names_exhausts = mk_case_names_exhausts descr' (map #1 new_dts);
- in
- add_datatype_def
- (config : datatype_config) new_type_names descr sorts types_syntax constr_syntax dt_info
- case_names_induct case_names_exhausts thy
- end;
-
-val add_datatype = gen_add_datatype cert_typ;
-val add_datatype_cmd = gen_add_datatype read_typ default_datatype_config;
-
-
-
-(** package setup **)
-
-(* setup theory *)
-
-val setup =
- DatatypeRepProofs.distinctness_limit_setup #>
- simproc_setup #>
- trfun_setup #>
- DatatypeInterpretation.init;
-
-
-(* outer syntax *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val datatype_decl =
- Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.binding -- P.opt_infix --
- (P.$$$ "=" |-- P.enum1 "|" (P.binding -- Scan.repeat P.typ -- P.opt_mixfix));
-
-fun mk_datatype args =
- let
- val names = map
- (fn ((((NONE, _), t), _), _) => Binding.name_of t | ((((SOME t, _), _), _), _) => t) args;
- val specs = map (fn ((((_, vs), t), mx), cons) =>
- (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
- in snd o add_datatype_cmd names specs end;
-
-val _ =
- OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl
- (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
-
-val _ =
- OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_goal
- (Scan.option (P.$$$ "(" |-- Scan.repeat1 P.name --| P.$$$ ")") -- Scan.repeat1 P.term
- >> (fn (alt_names, ts) => Toplevel.print
- o Toplevel.theory_to_proof (rep_datatype_cmd alt_names ts)));
-
-end;
-
-
-(* document antiquotation *)
-
-val _ = ThyOutput.antiquotation "datatype" Args.tyname
- (fn {source = src, context = ctxt, ...} => fn dtco =>
- let
- val thy = ProofContext.theory_of ctxt;
- val (vs, cos) = the_datatype_spec thy dtco;
- val ty = Type (dtco, map TFree vs);
- fun pretty_typ_bracket (ty as Type (_, _ :: _)) =
- Pretty.enclose "(" ")" [Syntax.pretty_typ ctxt ty]
- | pretty_typ_bracket ty =
- Syntax.pretty_typ ctxt ty;
- fun pretty_constr (co, tys) =
- (Pretty.block o Pretty.breaks)
- (Syntax.pretty_term ctxt (Const (co, tys ---> ty)) ::
- map pretty_typ_bracket tys);
- val pretty_datatype =
- Pretty.block
- (Pretty.command "datatype" :: Pretty.brk 1 ::
- Syntax.pretty_typ ctxt ty ::
- Pretty.str " =" :: Pretty.brk 1 ::
- flat (separate [Pretty.brk 1, Pretty.str "| "]
- (map (single o pretty_constr) cos)));
- in ThyOutput.output (ThyOutput.maybe_pretty_source (K pretty_datatype) src [()]) end);
-
-end;
-
--- a/src/HOL/Tools/datatype_package/datatype_realizer.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Tools/datatype_package/datatype_realizer.ML Fri Jun 19 17:23:21 2009 +0200
@@ -217,7 +217,7 @@
if ! Proofterm.proofs < 2 then thy
else let
val _ = message config "Adding realizers for induction and case analysis ..."
- val infos = map (DatatypePackage.the_datatype thy) names;
+ val infos = map (Datatype.the_datatype thy) names;
val info :: _ = infos;
in
thy
@@ -225,6 +225,6 @@
|> fold_rev (make_casedists (#sorts info)) infos
end;
-val setup = DatatypePackage.interpretation add_dt_realizers;
+val setup = Datatype.interpretation add_dt_realizers;
end;
--- a/src/HOL/Tools/datatype_package/datatype_rep_proofs.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Tools/datatype_package/datatype_rep_proofs.ML Fri Jun 19 17:23:21 2009 +0200
@@ -183,7 +183,7 @@
((1 upto (length constrs)) ~~ constrs)) (descr' ~~ rep_set_names');
val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy2) =
- InductivePackage.add_inductive_global (serial_string ())
+ Inductive.add_inductive_global (serial_string ())
{quiet_mode = #quiet config, verbose = false, kind = Thm.internalK,
alt_name = Binding.name big_rec_name, coind = false, no_elim = true, no_ind = false,
skip_mono = true, fork_mono = false}
@@ -195,7 +195,7 @@
val (typedefs, thy3) = thy2 |>
parent_path (#flat_names config) |>
fold_map (fn ((((name, mx), tvs), c), name') =>
- TypedefPackage.add_typedef false (SOME (Binding.name name')) (name, tvs, mx)
+ Typedef.add_typedef false (SOME (Binding.name name')) (name, tvs, mx)
(Collect $ Const (c, UnivT')) NONE
(rtac exI 1 THEN rtac CollectI 1 THEN
QUIET_BREADTH_FIRST (has_fewer_prems 1)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/function_package/fundef.ML Fri Jun 19 17:23:21 2009 +0200
@@ -0,0 +1,226 @@
+(* Title: HOL/Tools/function_package/fundef.ML
+ Author: Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions.
+Isar commands.
+*)
+
+signature FUNDEF =
+sig
+ val add_fundef : (binding * typ option * mixfix) list
+ -> (Attrib.binding * term) list
+ -> FundefCommon.fundef_config
+ -> local_theory
+ -> Proof.state
+ val add_fundef_cmd : (binding * string option * mixfix) list
+ -> (Attrib.binding * string) list
+ -> FundefCommon.fundef_config
+ -> local_theory
+ -> Proof.state
+
+ val termination_proof : term option -> local_theory -> Proof.state
+ val termination_proof_cmd : string option -> local_theory -> Proof.state
+ val termination : term option -> local_theory -> Proof.state
+ val termination_cmd : string option -> local_theory -> Proof.state
+
+ val setup : theory -> theory
+ val get_congs : Proof.context -> thm list
+end
+
+
+structure Fundef : FUNDEF =
+struct
+
+open FundefLib
+open FundefCommon
+
+val simp_attribs = map (Attrib.internal o K)
+ [Simplifier.simp_add,
+ Code.add_default_eqn_attribute,
+ Nitpick_Const_Simp_Thms.add,
+ Quickcheck_RecFun_Simp_Thms.add]
+
+val psimp_attribs = map (Attrib.internal o K)
+ [Simplifier.simp_add,
+ Nitpick_Const_Psimp_Thms.add]
+
+fun note_theorem ((name, atts), ths) =
+ LocalTheory.note Thm.generatedK ((Binding.qualified_name name, atts), ths)
+
+fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_"
+
+fun add_simps fnames post sort extra_qualify label moreatts simps lthy =
+ let
+ val spec = post simps
+ |> map (apfst (apsnd (fn ats => moreatts @ ats)))
+ |> map (apfst (apfst extra_qualify))
+
+ val (saved_spec_simps, lthy) =
+ fold_map (LocalTheory.note Thm.generatedK) spec lthy
+
+ val saved_simps = flat (map snd saved_spec_simps)
+ val simps_by_f = sort saved_simps
+
+ fun add_for_f fname simps =
+ note_theorem ((Long_Name.qualify fname label, []), simps) #> snd
+ in
+ (saved_simps,
+ fold2 add_for_f fnames simps_by_f lthy)
+ end
+
+fun gen_add_fundef is_external prep default_constraint fixspec eqns config lthy =
+ let
+ val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx))
+ val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy
+ val fixes = map (apfst (apfst Binding.name_of)) fixes0;
+ val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;
+ val (eqs, post, sort_cont, cnames) = FundefCommon.get_preproc lthy config ctxt' fixes spec
+
+ val defname = mk_defname fixes
+
+ val ((goalstate, cont), lthy) =
+ FundefMutual.prepare_fundef_mutual config defname fixes eqs lthy
+
+ fun afterqed [[proof]] lthy =
+ let
+ val FundefResult {fs, R, psimps, trsimps, simple_pinducts, termination,
+ domintros, cases, ...} =
+ cont (Thm.close_derivation proof)
+
+ val fnames = map (fst o fst) fixes
+ val qualify = Long_Name.qualify defname
+ val addsmps = add_simps fnames post sort_cont
+
+ val (((psimps', pinducts'), (_, [termination'])), lthy) =
+ lthy
+ |> addsmps (Binding.qualify false "partial") "psimps"
+ psimp_attribs psimps
+ ||> fold_option (snd oo addsmps I "simps" simp_attribs) trsimps
+ ||>> note_theorem ((qualify "pinduct",
+ [Attrib.internal (K (RuleCases.case_names cnames)),
+ Attrib.internal (K (RuleCases.consumes 1)),
+ Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts)
+ ||>> note_theorem ((qualify "termination", []), [termination])
+ ||> (snd o note_theorem ((qualify "cases",
+ [Attrib.internal (K (RuleCases.case_names cnames))]), [cases]))
+ ||> fold_option (snd oo curry note_theorem (qualify "domintros", [])) domintros
+
+ val cdata = FundefCtxData { add_simps=addsmps, case_names=cnames, psimps=psimps',
+ pinducts=snd pinducts', termination=termination',
+ fs=fs, R=R, defname=defname }
+ val _ =
+ if not is_external then ()
+ else Specification.print_consts lthy (K false) (map fst fixes)
+ in
+ lthy
+ |> LocalTheory.declaration (add_fundef_data o morph_fundef_data cdata)
+ end
+ in
+ lthy
+ |> is_external ? LocalTheory.set_group (serial_string ())
+ |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
+ |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd
+ end
+
+val add_fundef = gen_add_fundef false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
+val add_fundef_cmd = gen_add_fundef true Specification.read_spec "_::type"
+
+fun gen_termination_proof prep_term raw_term_opt lthy =
+ let
+ val term_opt = Option.map (prep_term lthy) raw_term_opt
+ val data = the (case term_opt of
+ SOME t => (import_fundef_data t lthy
+ handle Option.Option =>
+ error ("Not a function: " ^ quote (Syntax.string_of_term lthy t)))
+ | NONE => (import_last_fundef lthy handle Option.Option => error "Not a function"))
+
+ val FundefCtxData { termination, R, add_simps, case_names, psimps,
+ pinducts, defname, ...} = data
+ val domT = domain_type (fastype_of R)
+ val goal = HOLogic.mk_Trueprop
+ (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
+ fun afterqed [[totality]] lthy =
+ let
+ val totality = Thm.close_derivation totality
+ val remove_domain_condition =
+ full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals])
+ val tsimps = map remove_domain_condition psimps
+ val tinduct = map remove_domain_condition pinducts
+ val qualify = Long_Name.qualify defname;
+ in
+ lthy
+ |> add_simps I "simps" simp_attribs tsimps |> snd
+ |> note_theorem
+ ((qualify "induct",
+ [Attrib.internal (K (RuleCases.case_names case_names))]),
+ tinduct) |> snd
+ end
+ in
+ lthy
+ |> ProofContext.note_thmss ""
+ [((Binding.empty, [ContextRules.rule_del]), [([allI], [])])] |> snd
+ |> ProofContext.note_thmss ""
+ [((Binding.empty, [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
+ |> ProofContext.note_thmss ""
+ [((Binding.name "termination", [ContextRules.intro_bang (SOME 0)]),
+ [([Goal.norm_result termination], [])])] |> snd
+ |> Proof.theorem_i NONE afterqed [[(goal, [])]]
+ end
+
+val termination_proof = gen_termination_proof Syntax.check_term;
+val termination_proof_cmd = gen_termination_proof Syntax.read_term;
+
+fun termination term_opt lthy =
+ lthy
+ |> LocalTheory.set_group (serial_string ())
+ |> termination_proof term_opt;
+
+fun termination_cmd term_opt lthy =
+ lthy
+ |> LocalTheory.set_group (serial_string ())
+ |> termination_proof_cmd term_opt;
+
+
+(* Datatype hook to declare datatype congs as "fundef_congs" *)
+
+
+fun add_case_cong n thy =
+ Context.theory_map (FundefCtxTree.map_fundef_congs (Thm.add_thm
+ (Datatype.get_datatype thy n |> the
+ |> #case_cong
+ |> safe_mk_meta_eq)))
+ thy
+
+val setup_case_cong = Datatype.interpretation (K (fold add_case_cong))
+
+
+(* setup *)
+
+val setup =
+ Attrib.setup @{binding fundef_cong}
+ (Attrib.add_del FundefCtxTree.cong_add FundefCtxTree.cong_del)
+ "declaration of congruence rule for function definitions"
+ #> setup_case_cong
+ #> FundefRelation.setup
+ #> FundefCommon.TerminationSimps.setup
+
+val get_congs = FundefCtxTree.get_fundef_congs
+
+
+(* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val _ =
+ OuterSyntax.local_theory_to_proof "function" "define general recursive functions" K.thy_goal
+ (fundef_parser default_config
+ >> (fn ((config, fixes), statements) => add_fundef_cmd fixes statements config));
+
+val _ =
+ OuterSyntax.local_theory_to_proof "termination" "prove termination of a recursive function" K.thy_goal
+ (Scan.option P.term >> termination_cmd);
+
+end;
+
+
+end
--- a/src/HOL/Tools/function_package/fundef_datatype.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Tools/function_package/fundef_datatype.ML Fri Jun 19 17:23:21 2009 +0200
@@ -40,7 +40,7 @@
let
val (hd, args) = strip_comb t
in
- (((case DatatypePackage.datatype_of_constr thy (fst (dest_Const hd)) of
+ (((case Datatype.datatype_of_constr thy (fst (dest_Const hd)) of
SOME _ => ()
| NONE => err "Non-constructor pattern")
handle TERM ("dest_Const", _) => err "Non-constructor patterns");
@@ -103,7 +103,7 @@
fun inst_constrs_of thy (T as Type (name, _)) =
map (fn (Cn,CT) => Envir.subst_TVars (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT)))
- (the (DatatypePackage.get_datatype_constrs thy name))
+ (the (Datatype.get_datatype_constrs thy name))
| inst_constrs_of thy _ = raise Match
@@ -144,7 +144,7 @@
let
val T = fastype_of v
val (tname, _) = dest_Type T
- val {exhaustion=case_thm, ...} = DatatypePackage.the_datatype thy tname
+ val {exhaustion=case_thm, ...} = Datatype.the_datatype thy tname
val constrs = inst_constrs_of thy T
val c_cases = map (constr_case thy P idx (v :: vs) pts) constrs
in
@@ -211,7 +211,7 @@
SOME (Method.Source_i (Args.src (("HOL.auto", []), Position.none))))
fun termination_by method int =
- FundefPackage.termination_proof NONE
+ Fundef.termination_proof NONE
#> Proof.global_future_terminal_proof
(Method.Basic (method, Position.none), NONE) int
@@ -313,8 +313,8 @@
|> termination_by (FundefCommon.get_termination_prover lthy) int
end;
-val add_fun = gen_fun FundefPackage.add_fundef
-val add_fun_cmd = gen_fun FundefPackage.add_fundef_cmd
+val add_fun = gen_fun Fundef.add_fundef
+val add_fun_cmd = gen_fun Fundef.add_fundef_cmd
--- a/src/HOL/Tools/function_package/fundef_package.ML Thu Jun 18 18:31:14 2009 -0700
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,226 +0,0 @@
-(* Title: HOL/Tools/function_package/fundef_package.ML
- Author: Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-Isar commands.
-*)
-
-signature FUNDEF_PACKAGE =
-sig
- val add_fundef : (binding * typ option * mixfix) list
- -> (Attrib.binding * term) list
- -> FundefCommon.fundef_config
- -> local_theory
- -> Proof.state
- val add_fundef_cmd : (binding * string option * mixfix) list
- -> (Attrib.binding * string) list
- -> FundefCommon.fundef_config
- -> local_theory
- -> Proof.state
-
- val termination_proof : term option -> local_theory -> Proof.state
- val termination_proof_cmd : string option -> local_theory -> Proof.state
- val termination : term option -> local_theory -> Proof.state
- val termination_cmd : string option -> local_theory -> Proof.state
-
- val setup : theory -> theory
- val get_congs : Proof.context -> thm list
-end
-
-
-structure FundefPackage : FUNDEF_PACKAGE =
-struct
-
-open FundefLib
-open FundefCommon
-
-val simp_attribs = map (Attrib.internal o K)
- [Simplifier.simp_add,
- Code.add_default_eqn_attribute,
- Nitpick_Const_Simp_Thms.add,
- Quickcheck_RecFun_Simp_Thms.add]
-
-val psimp_attribs = map (Attrib.internal o K)
- [Simplifier.simp_add,
- Nitpick_Const_Psimp_Thms.add]
-
-fun note_theorem ((name, atts), ths) =
- LocalTheory.note Thm.generatedK ((Binding.qualified_name name, atts), ths)
-
-fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_"
-
-fun add_simps fnames post sort extra_qualify label moreatts simps lthy =
- let
- val spec = post simps
- |> map (apfst (apsnd (fn ats => moreatts @ ats)))
- |> map (apfst (apfst extra_qualify))
-
- val (saved_spec_simps, lthy) =
- fold_map (LocalTheory.note Thm.generatedK) spec lthy
-
- val saved_simps = flat (map snd saved_spec_simps)
- val simps_by_f = sort saved_simps
-
- fun add_for_f fname simps =
- note_theorem ((Long_Name.qualify fname label, []), simps) #> snd
- in
- (saved_simps,
- fold2 add_for_f fnames simps_by_f lthy)
- end
-
-fun gen_add_fundef is_external prep default_constraint fixspec eqns config lthy =
- let
- val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx))
- val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy
- val fixes = map (apfst (apfst Binding.name_of)) fixes0;
- val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;
- val (eqs, post, sort_cont, cnames) = FundefCommon.get_preproc lthy config ctxt' fixes spec
-
- val defname = mk_defname fixes
-
- val ((goalstate, cont), lthy) =
- FundefMutual.prepare_fundef_mutual config defname fixes eqs lthy
-
- fun afterqed [[proof]] lthy =
- let
- val FundefResult {fs, R, psimps, trsimps, simple_pinducts, termination,
- domintros, cases, ...} =
- cont (Thm.close_derivation proof)
-
- val fnames = map (fst o fst) fixes
- val qualify = Long_Name.qualify defname
- val addsmps = add_simps fnames post sort_cont
-
- val (((psimps', pinducts'), (_, [termination'])), lthy) =
- lthy
- |> addsmps (Binding.qualify false "partial") "psimps"
- psimp_attribs psimps
- ||> fold_option (snd oo addsmps I "simps" simp_attribs) trsimps
- ||>> note_theorem ((qualify "pinduct",
- [Attrib.internal (K (RuleCases.case_names cnames)),
- Attrib.internal (K (RuleCases.consumes 1)),
- Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts)
- ||>> note_theorem ((qualify "termination", []), [termination])
- ||> (snd o note_theorem ((qualify "cases",
- [Attrib.internal (K (RuleCases.case_names cnames))]), [cases]))
- ||> fold_option (snd oo curry note_theorem (qualify "domintros", [])) domintros
-
- val cdata = FundefCtxData { add_simps=addsmps, case_names=cnames, psimps=psimps',
- pinducts=snd pinducts', termination=termination',
- fs=fs, R=R, defname=defname }
- val _ =
- if not is_external then ()
- else Specification.print_consts lthy (K false) (map fst fixes)
- in
- lthy
- |> LocalTheory.declaration (add_fundef_data o morph_fundef_data cdata)
- end
- in
- lthy
- |> is_external ? LocalTheory.set_group (serial_string ())
- |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
- |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd
- end
-
-val add_fundef = gen_add_fundef false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
-val add_fundef_cmd = gen_add_fundef true Specification.read_spec "_::type"
-
-fun gen_termination_proof prep_term raw_term_opt lthy =
- let
- val term_opt = Option.map (prep_term lthy) raw_term_opt
- val data = the (case term_opt of
- SOME t => (import_fundef_data t lthy
- handle Option.Option =>
- error ("Not a function: " ^ quote (Syntax.string_of_term lthy t)))
- | NONE => (import_last_fundef lthy handle Option.Option => error "Not a function"))
-
- val FundefCtxData { termination, R, add_simps, case_names, psimps,
- pinducts, defname, ...} = data
- val domT = domain_type (fastype_of R)
- val goal = HOLogic.mk_Trueprop
- (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
- fun afterqed [[totality]] lthy =
- let
- val totality = Thm.close_derivation totality
- val remove_domain_condition =
- full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals])
- val tsimps = map remove_domain_condition psimps
- val tinduct = map remove_domain_condition pinducts
- val qualify = Long_Name.qualify defname;
- in
- lthy
- |> add_simps I "simps" simp_attribs tsimps |> snd
- |> note_theorem
- ((qualify "induct",
- [Attrib.internal (K (RuleCases.case_names case_names))]),
- tinduct) |> snd
- end
- in
- lthy
- |> ProofContext.note_thmss ""
- [((Binding.empty, [ContextRules.rule_del]), [([allI], [])])] |> snd
- |> ProofContext.note_thmss ""
- [((Binding.empty, [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
- |> ProofContext.note_thmss ""
- [((Binding.name "termination", [ContextRules.intro_bang (SOME 0)]),
- [([Goal.norm_result termination], [])])] |> snd
- |> Proof.theorem_i NONE afterqed [[(goal, [])]]
- end
-
-val termination_proof = gen_termination_proof Syntax.check_term;
-val termination_proof_cmd = gen_termination_proof Syntax.read_term;
-
-fun termination term_opt lthy =
- lthy
- |> LocalTheory.set_group (serial_string ())
- |> termination_proof term_opt;
-
-fun termination_cmd term_opt lthy =
- lthy
- |> LocalTheory.set_group (serial_string ())
- |> termination_proof_cmd term_opt;
-
-
-(* Datatype hook to declare datatype congs as "fundef_congs" *)
-
-
-fun add_case_cong n thy =
- Context.theory_map (FundefCtxTree.map_fundef_congs (Thm.add_thm
- (DatatypePackage.get_datatype thy n |> the
- |> #case_cong
- |> safe_mk_meta_eq)))
- thy
-
-val setup_case_cong = DatatypePackage.interpretation (K (fold add_case_cong))
-
-
-(* setup *)
-
-val setup =
- Attrib.setup @{binding fundef_cong}
- (Attrib.add_del FundefCtxTree.cong_add FundefCtxTree.cong_del)
- "declaration of congruence rule for function definitions"
- #> setup_case_cong
- #> FundefRelation.setup
- #> FundefCommon.TerminationSimps.setup
-
-val get_congs = FundefCtxTree.get_fundef_congs
-
-
-(* outer syntax *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ =
- OuterSyntax.local_theory_to_proof "function" "define general recursive functions" K.thy_goal
- (fundef_parser default_config
- >> (fn ((config, fixes), statements) => add_fundef_cmd fixes statements config));
-
-val _ =
- OuterSyntax.local_theory_to_proof "termination" "prove termination of a recursive function" K.thy_goal
- (Scan.option P.term >> termination_cmd);
-
-end;
-
-
-end
--- a/src/HOL/Tools/function_package/inductive_wrap.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Tools/function_package/inductive_wrap.ML Fri Jun 19 17:23:21 2009 +0200
@@ -1,5 +1,4 @@
(* Title: HOL/Tools/function_package/inductive_wrap.ML
- ID: $Id$
Author: Alexander Krauss, TU Muenchen
@@ -40,7 +39,7 @@
fun inductive_def defs (((R, T), mixfix), lthy) =
let
val ({intrs = intrs_gen, elims = [elim_gen], preds = [ Rdef ], induct, ...}, lthy) =
- InductivePackage.add_inductive_i
+ Inductive.add_inductive_i
{quiet_mode = false,
verbose = ! Toplevel.debug,
kind = Thm.internalK,
--- a/src/HOL/Tools/function_package/pattern_split.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Tools/function_package/pattern_split.ML Fri Jun 19 17:23:21 2009 +0200
@@ -41,7 +41,7 @@
(* This is copied from "fundef_datatype.ML" *)
fun inst_constrs_of thy (T as Type (name, _)) =
map (fn (Cn,CT) => Envir.subst_TVars (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT)))
- (the (DatatypePackage.get_datatype_constrs thy name))
+ (the (Datatype.get_datatype_constrs thy name))
| inst_constrs_of thy T = raise TYPE ("inst_constrs_of", [T], [])
--- a/src/HOL/Tools/function_package/size.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Tools/function_package/size.ML Fri Jun 19 17:23:21 2009 +0200
@@ -44,14 +44,14 @@
| SOME t => t);
fun is_poly thy (DtType (name, dts)) =
- (case DatatypePackage.get_datatype thy name of
+ (case Datatype.get_datatype thy name of
NONE => false
| SOME _ => exists (is_poly thy) dts)
| is_poly _ _ = true;
fun constrs_of thy name =
let
- val {descr, index, ...} = DatatypePackage.the_datatype thy name
+ val {descr, index, ...} = Datatype.the_datatype thy name
val SOME (_, _, constrs) = AList.lookup op = descr index
in constrs end;
@@ -220,7 +220,7 @@
fun add_size_thms config (new_type_names as name :: _) thy =
let
- val info as {descr, alt_names, ...} = DatatypePackage.the_datatype thy name;
+ val info as {descr, alt_names, ...} = Datatype.the_datatype thy name;
val prefix = Long_Name.map_base_name (K (space_implode "_"
(the_default (map Long_Name.base_name new_type_names) alt_names))) name;
val no_size = exists (fn (_, (_, _, constrs)) => exists (fn (_, cargs) => exists (fn dt =>
@@ -237,6 +237,6 @@
val size_thms = snd oo (the oo lookup_size);
-val setup = DatatypePackage.interpretation add_size_thms;
+val setup = Datatype.interpretation add_size_thms;
end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/inductive.ML Fri Jun 19 17:23:21 2009 +0200
@@ -0,0 +1,968 @@
+(* Title: HOL/Tools/inductive.ML
+ Author: Lawrence C Paulson, Cambridge University Computer Laboratory
+ Author: Stefan Berghofer and Markus Wenzel, TU Muenchen
+
+(Co)Inductive Definition module for HOL.
+
+Features:
+ * least or greatest fixedpoints
+ * mutually recursive definitions
+ * definitions involving arbitrary monotone operators
+ * automatically proves introduction and elimination rules
+
+ Introduction rules have the form
+ [| M Pj ti, ..., Q x, ... |] ==> Pk t
+ where M is some monotone operator (usually the identity)
+ Q x is any side condition on the free variables
+ ti, t are any terms
+ Pj, Pk are two of the predicates being defined in mutual recursion
+*)
+
+signature BASIC_INDUCTIVE =
+sig
+ type inductive_result
+ val morph_result: morphism -> inductive_result -> inductive_result
+ type inductive_info
+ val the_inductive: Proof.context -> string -> inductive_info
+ val print_inductives: Proof.context -> unit
+ val mono_add: attribute
+ val mono_del: attribute
+ val get_monos: Proof.context -> thm list
+ val mk_cases: Proof.context -> term -> thm
+ val inductive_forall_name: string
+ val inductive_forall_def: thm
+ val rulify: thm -> thm
+ val inductive_cases: (Attrib.binding * string list) list -> local_theory ->
+ thm list list * local_theory
+ val inductive_cases_i: (Attrib.binding * term list) list -> local_theory ->
+ thm list list * local_theory
+ type inductive_flags
+ val add_inductive_i:
+ inductive_flags -> ((binding * typ) * mixfix) list ->
+ (string * typ) list -> (Attrib.binding * term) list -> thm list -> local_theory ->
+ inductive_result * local_theory
+ val add_inductive: bool -> bool ->
+ (binding * string option * mixfix) list ->
+ (binding * string option * mixfix) list ->
+ (Attrib.binding * string) list ->
+ (Facts.ref * Attrib.src list) list ->
+ bool -> local_theory -> inductive_result * local_theory
+ val add_inductive_global: string -> inductive_flags ->
+ ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
+ thm list -> theory -> inductive_result * theory
+ val arities_of: thm -> (string * int) list
+ val params_of: thm -> term list
+ val partition_rules: thm -> thm list -> (string * thm list) list
+ val partition_rules': thm -> (thm * 'a) list -> (string * (thm * 'a) list) list
+ val unpartition_rules: thm list -> (string * 'a list) list -> 'a list
+ val infer_intro_vars: thm -> int -> thm list -> term list list
+ val setup: theory -> theory
+end;
+
+signature INDUCTIVE =
+sig
+ include BASIC_INDUCTIVE
+ type add_ind_def
+ val declare_rules: string -> binding -> bool -> bool -> string list ->
+ thm list -> binding list -> Attrib.src list list -> (thm * string list) list ->
+ thm -> local_theory -> thm list * thm list * thm * local_theory
+ val add_ind_def: add_ind_def
+ val gen_add_inductive_i: add_ind_def -> inductive_flags ->
+ ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
+ thm list -> local_theory -> inductive_result * local_theory
+ val gen_add_inductive: add_ind_def -> bool -> bool ->
+ (binding * string option * mixfix) list ->
+ (binding * string option * mixfix) list ->
+ (Attrib.binding * string) list -> (Facts.ref * Attrib.src list) list ->
+ bool -> local_theory -> inductive_result * local_theory
+ val gen_ind_decl: add_ind_def -> bool ->
+ OuterParse.token list -> (bool -> local_theory -> local_theory) * OuterParse.token list
+end;
+
+structure Inductive: INDUCTIVE =
+struct
+
+
+(** theory context references **)
+
+val inductive_forall_name = "HOL.induct_forall";
+val inductive_forall_def = thm "induct_forall_def";
+val inductive_conj_name = "HOL.induct_conj";
+val inductive_conj_def = thm "induct_conj_def";
+val inductive_conj = thms "induct_conj";
+val inductive_atomize = thms "induct_atomize";
+val inductive_rulify = thms "induct_rulify";
+val inductive_rulify_fallback = thms "induct_rulify_fallback";
+
+val notTrueE = TrueI RSN (2, notE);
+val notFalseI = Seq.hd (atac 1 notI);
+val simp_thms' = map (fn s => mk_meta_eq (the (find_first
+ (equal (OldGoals.read_prop @{theory HOL} s) o prop_of) simp_thms)))
+ ["(~True) = False", "(~False) = True",
+ "(True --> ?P) = ?P", "(False --> ?P) = True",
+ "(?P & True) = ?P", "(True & ?P) = ?P"];
+
+
+
+(** context data **)
+
+type inductive_result =
+ {preds: term list, elims: thm list, raw_induct: thm,
+ induct: thm, intrs: thm list};
+
+fun morph_result phi {preds, elims, raw_induct: thm, induct, intrs} =
+ let
+ val term = Morphism.term phi;
+ val thm = Morphism.thm phi;
+ val fact = Morphism.fact phi;
+ in
+ {preds = map term preds, elims = fact elims, raw_induct = thm raw_induct,
+ induct = thm induct, intrs = fact intrs}
+ end;
+
+type inductive_info =
+ {names: string list, coind: bool} * inductive_result;
+
+structure InductiveData = GenericDataFun
+(
+ type T = inductive_info Symtab.table * thm list;
+ val empty = (Symtab.empty, []);
+ val extend = I;
+ fun merge _ ((tab1, monos1), (tab2, monos2)) =
+ (Symtab.merge (K true) (tab1, tab2), Thm.merge_thms (monos1, monos2));
+);
+
+val get_inductives = InductiveData.get o Context.Proof;
+
+fun print_inductives ctxt =
+ let
+ val (tab, monos) = get_inductives ctxt;
+ val space = Consts.space_of (ProofContext.consts_of ctxt);
+ in
+ [Pretty.strs ("(co)inductives:" :: map #1 (NameSpace.extern_table (space, tab))),
+ Pretty.big_list "monotonicity rules:" (map (ProofContext.pretty_thm ctxt) monos)]
+ |> Pretty.chunks |> Pretty.writeln
+ end;
+
+
+(* get and put data *)
+
+fun the_inductive ctxt name =
+ (case Symtab.lookup (#1 (get_inductives ctxt)) name of
+ NONE => error ("Unknown (co)inductive predicate " ^ quote name)
+ | SOME info => info);
+
+fun put_inductives names info = InductiveData.map
+ (apfst (fold (fn name => Symtab.update (name, info)) names));
+
+
+
+(** monotonicity rules **)
+
+val get_monos = #2 o get_inductives;
+val map_monos = InductiveData.map o apsnd;
+
+fun mk_mono thm =
+ let
+ val concl = concl_of thm;
+ fun eq2mono thm' = [thm' RS (thm' RS eq_to_mono)] @
+ (case concl of
+ (_ $ (_ $ (Const ("Not", _) $ _) $ _)) => []
+ | _ => [thm' RS (thm' RS eq_to_mono2)]);
+ fun dest_less_concl thm = dest_less_concl (thm RS le_funD)
+ handle THM _ => thm RS le_boolD
+ in
+ case concl of
+ Const ("==", _) $ _ $ _ => eq2mono (thm RS meta_eq_to_obj_eq)
+ | _ $ (Const ("op =", _) $ _ $ _) => eq2mono thm
+ | _ $ (Const ("HOL.ord_class.less_eq", _) $ _ $ _) =>
+ [dest_less_concl (Seq.hd (REPEAT (FIRSTGOAL
+ (resolve_tac [le_funI, le_boolI'])) thm))]
+ | _ => [thm]
+ end handle THM _ => error ("Bad monotonicity theorem:\n" ^ Display.string_of_thm thm);
+
+val mono_add = Thm.declaration_attribute (map_monos o fold Thm.add_thm o mk_mono);
+val mono_del = Thm.declaration_attribute (map_monos o fold Thm.del_thm o mk_mono);
+
+
+
+(** misc utilities **)
+
+fun message quiet_mode s = if quiet_mode then () else writeln s;
+fun clean_message quiet_mode s = if ! quick_and_dirty then () else message quiet_mode s;
+
+fun coind_prefix true = "co"
+ | coind_prefix false = "";
+
+fun log (b:int) m n = if m >= n then 0 else 1 + log b (b * m) n;
+
+fun make_bool_args f g [] i = []
+ | make_bool_args f g (x :: xs) i =
+ (if i mod 2 = 0 then f x else g x) :: make_bool_args f g xs (i div 2);
+
+fun make_bool_args' xs =
+ make_bool_args (K HOLogic.false_const) (K HOLogic.true_const) xs;
+
+fun find_arg T x [] = sys_error "find_arg"
+ | find_arg T x ((p as (_, (SOME _, _))) :: ps) =
+ apsnd (cons p) (find_arg T x ps)
+ | find_arg T x ((p as (U, (NONE, y))) :: ps) =
+ if (T: typ) = U then (y, (U, (SOME x, y)) :: ps)
+ else apsnd (cons p) (find_arg T x ps);
+
+fun make_args Ts xs =
+ map (fn (T, (NONE, ())) => Const (@{const_name undefined}, T) | (_, (SOME t, ())) => t)
+ (fold (fn (t, T) => snd o find_arg T t) xs (map (rpair (NONE, ())) Ts));
+
+fun make_args' Ts xs Us =
+ fst (fold_map (fn T => find_arg T ()) Us (Ts ~~ map (pair NONE) xs));
+
+fun dest_predicate cs params t =
+ let
+ val k = length params;
+ val (c, ts) = strip_comb t;
+ val (xs, ys) = chop k ts;
+ val i = find_index_eq c cs;
+ in
+ if xs = params andalso i >= 0 then
+ SOME (c, i, ys, chop (length ys)
+ (List.drop (binder_types (fastype_of c), k)))
+ else NONE
+ end;
+
+fun mk_names a 0 = []
+ | mk_names a 1 = [a]
+ | mk_names a n = map (fn i => a ^ string_of_int i) (1 upto n);
+
+
+
+(** process rules **)
+
+local
+
+fun err_in_rule ctxt name t msg =
+ error (cat_lines ["Ill-formed introduction rule " ^ quote name,
+ Syntax.string_of_term ctxt t, msg]);
+
+fun err_in_prem ctxt name t p msg =
+ error (cat_lines ["Ill-formed premise", Syntax.string_of_term ctxt p,
+ "in introduction rule " ^ quote name, Syntax.string_of_term ctxt t, msg]);
+
+val bad_concl = "Conclusion of introduction rule must be an inductive predicate";
+
+val bad_ind_occ = "Inductive predicate occurs in argument of inductive predicate";
+
+val bad_app = "Inductive predicate must be applied to parameter(s) ";
+
+fun atomize_term thy = MetaSimplifier.rewrite_term thy inductive_atomize [];
+
+in
+
+fun check_rule ctxt cs params ((binding, att), rule) =
+ let
+ val err_name = Binding.str_of binding;
+ val params' = Term.variant_frees rule (Logic.strip_params rule);
+ val frees = rev (map Free params');
+ val concl = subst_bounds (frees, Logic.strip_assums_concl rule);
+ val prems = map (curry subst_bounds frees) (Logic.strip_assums_hyp rule);
+ val rule' = Logic.list_implies (prems, concl);
+ val aprems = map (atomize_term (ProofContext.theory_of ctxt)) prems;
+ val arule = list_all_free (params', Logic.list_implies (aprems, concl));
+
+ fun check_ind err t = case dest_predicate cs params t of
+ NONE => err (bad_app ^
+ commas (map (Syntax.string_of_term ctxt) params))
+ | SOME (_, _, ys, _) =>
+ if exists (fn c => exists (fn t => Logic.occs (c, t)) ys) cs
+ then err bad_ind_occ else ();
+
+ fun check_prem' prem t =
+ if head_of t mem cs then
+ check_ind (err_in_prem ctxt err_name rule prem) t
+ else (case t of
+ Abs (_, _, t) => check_prem' prem t
+ | t $ u => (check_prem' prem t; check_prem' prem u)
+ | _ => ());
+
+ fun check_prem (prem, aprem) =
+ if can HOLogic.dest_Trueprop aprem then check_prem' prem prem
+ else err_in_prem ctxt err_name rule prem "Non-atomic premise";
+ in
+ (case concl of
+ Const ("Trueprop", _) $ t =>
+ if head_of t mem cs then
+ (check_ind (err_in_rule ctxt err_name rule') t;
+ List.app check_prem (prems ~~ aprems))
+ else err_in_rule ctxt err_name rule' bad_concl
+ | _ => err_in_rule ctxt err_name rule' bad_concl);
+ ((binding, att), arule)
+ end;
+
+val rulify =
+ hol_simplify inductive_conj
+ #> hol_simplify inductive_rulify
+ #> hol_simplify inductive_rulify_fallback
+ #> Simplifier.norm_hhf;
+
+end;
+
+
+
+(** proofs for (co)inductive predicates **)
+
+(* prove monotonicity *)
+
+fun prove_mono quiet_mode skip_mono fork_mono predT fp_fun monos ctxt =
+ (message (quiet_mode orelse skip_mono andalso !quick_and_dirty orelse fork_mono)
+ " Proving monotonicity ...";
+ (if skip_mono then SkipProof.prove else if fork_mono then Goal.prove_future else Goal.prove) ctxt
+ [] []
+ (HOLogic.mk_Trueprop
+ (Const (@{const_name Orderings.mono}, (predT --> predT) --> HOLogic.boolT) $ fp_fun))
+ (fn _ => EVERY [rtac @{thm monoI} 1,
+ REPEAT (resolve_tac [le_funI, le_boolI'] 1),
+ REPEAT (FIRST
+ [atac 1,
+ resolve_tac (List.concat (map mk_mono monos) @ get_monos ctxt) 1,
+ etac le_funE 1, dtac le_boolD 1])]));
+
+
+(* prove introduction rules *)
+
+fun prove_intrs quiet_mode coind mono fp_def k params intr_ts rec_preds_defs ctxt =
+ let
+ val _ = clean_message quiet_mode " Proving the introduction rules ...";
+
+ val unfold = funpow k (fn th => th RS fun_cong)
+ (mono RS (fp_def RS
+ (if coind then def_gfp_unfold else def_lfp_unfold)));
+
+ fun select_disj 1 1 = []
+ | select_disj _ 1 = [rtac disjI1]
+ | select_disj n i = (rtac disjI2)::(select_disj (n - 1) (i - 1));
+
+ val rules = [refl, TrueI, notFalseI, exI, conjI];
+
+ val intrs = map_index (fn (i, intr) => rulify
+ (SkipProof.prove ctxt (map (fst o dest_Free) params) [] intr (fn _ => EVERY
+ [rewrite_goals_tac rec_preds_defs,
+ rtac (unfold RS iffD2) 1,
+ EVERY1 (select_disj (length intr_ts) (i + 1)),
+ (*Not ares_tac, since refl must be tried before any equality assumptions;
+ backtracking may occur if the premises have extra variables!*)
+ DEPTH_SOLVE_1 (resolve_tac rules 1 APPEND assume_tac 1)]))) intr_ts
+
+ in (intrs, unfold) end;
+
+
+(* prove elimination rules *)
+
+fun prove_elims quiet_mode cs params intr_ts intr_names unfold rec_preds_defs ctxt =
+ let
+ val _ = clean_message quiet_mode " Proving the elimination rules ...";
+
+ val ([pname], ctxt') = ctxt |>
+ Variable.add_fixes (map (fst o dest_Free) params) |> snd |>
+ Variable.variant_fixes ["P"];
+ val P = HOLogic.mk_Trueprop (Free (pname, HOLogic.boolT));
+
+ fun dest_intr r =
+ (the (dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r))),
+ Logic.strip_assums_hyp r, Logic.strip_params r);
+
+ val intrs = map dest_intr intr_ts ~~ intr_names;
+
+ val rules1 = [disjE, exE, FalseE];
+ val rules2 = [conjE, FalseE, notTrueE];
+
+ fun prove_elim c =
+ let
+ val Ts = List.drop (binder_types (fastype_of c), length params);
+ val (anames, ctxt'') = Variable.variant_fixes (mk_names "a" (length Ts)) ctxt';
+ val frees = map Free (anames ~~ Ts);
+
+ fun mk_elim_prem ((_, _, us, _), ts, params') =
+ list_all (params',
+ Logic.list_implies (map (HOLogic.mk_Trueprop o HOLogic.mk_eq)
+ (frees ~~ us) @ ts, P));
+ val c_intrs = (List.filter (equal c o #1 o #1 o #1) intrs);
+ val prems = HOLogic.mk_Trueprop (list_comb (c, params @ frees)) ::
+ map mk_elim_prem (map #1 c_intrs)
+ in
+ (SkipProof.prove ctxt'' [] prems P
+ (fn {prems, ...} => EVERY
+ [cut_facts_tac [hd prems] 1,
+ rewrite_goals_tac rec_preds_defs,
+ dtac (unfold RS iffD1) 1,
+ REPEAT (FIRSTGOAL (eresolve_tac rules1)),
+ REPEAT (FIRSTGOAL (eresolve_tac rules2)),
+ EVERY (map (fn prem =>
+ DEPTH_SOLVE_1 (ares_tac [rewrite_rule rec_preds_defs prem, conjI] 1)) (tl prems))])
+ |> rulify
+ |> singleton (ProofContext.export ctxt'' ctxt),
+ map #2 c_intrs)
+ end
+
+ in map prove_elim cs end;
+
+
+(* derivation of simplified elimination rules *)
+
+local
+
+(*delete needless equality assumptions*)
+val refl_thin = Goal.prove_global @{theory HOL} [] [] @{prop "!!P. a = a ==> P ==> P"}
+ (fn _ => assume_tac 1);
+val elim_rls = [asm_rl, FalseE, refl_thin, conjE, exE];
+val elim_tac = REPEAT o Tactic.eresolve_tac elim_rls;
+
+fun simp_case_tac ss i =
+ EVERY' [elim_tac, asm_full_simp_tac ss, elim_tac, REPEAT o bound_hyp_subst_tac] i;
+
+in
+
+fun mk_cases ctxt prop =
+ let
+ val thy = ProofContext.theory_of ctxt;
+ val ss = Simplifier.local_simpset_of ctxt;
+
+ fun err msg =
+ error (Pretty.string_of (Pretty.block
+ [Pretty.str msg, Pretty.fbrk, Syntax.pretty_term ctxt prop]));
+
+ val elims = Induct.find_casesP ctxt prop;
+
+ val cprop = Thm.cterm_of thy prop;
+ val tac = ALLGOALS (simp_case_tac ss) THEN prune_params_tac;
+ fun mk_elim rl =
+ Thm.implies_intr cprop (Tactic.rule_by_tactic tac (Thm.assume cprop RS rl))
+ |> singleton (Variable.export (Variable.auto_fixes prop ctxt) ctxt);
+ in
+ (case get_first (try mk_elim) elims of
+ SOME r => r
+ | NONE => err "Proposition not an inductive predicate:")
+ end;
+
+end;
+
+
+(* inductive_cases *)
+
+fun gen_inductive_cases prep_att prep_prop args lthy =
+ let
+ val thy = ProofContext.theory_of lthy;
+ val facts = args |> map (fn ((a, atts), props) =>
+ ((a, map (prep_att thy) atts),
+ map (Thm.no_attributes o single o mk_cases lthy o prep_prop lthy) props));
+ in lthy |> LocalTheory.notes Thm.generatedK facts |>> map snd end;
+
+val inductive_cases = gen_inductive_cases Attrib.intern_src Syntax.read_prop;
+val inductive_cases_i = gen_inductive_cases (K I) Syntax.check_prop;
+
+
+val ind_cases_setup =
+ Method.setup @{binding ind_cases}
+ (Scan.lift (Scan.repeat1 Args.name_source --
+ Scan.optional (Args.$$$ "for" |-- Scan.repeat1 Args.name) []) >>
+ (fn (raw_props, fixes) => fn ctxt =>
+ let
+ val (_, ctxt') = Variable.add_fixes fixes ctxt;
+ val props = Syntax.read_props ctxt' raw_props;
+ val ctxt'' = fold Variable.declare_term props ctxt';
+ val rules = ProofContext.export ctxt'' ctxt (map (mk_cases ctxt'') props)
+ in Method.erule 0 rules end))
+ "dynamic case analysis on predicates";
+
+
+(* prove induction rule *)
+
+fun prove_indrule quiet_mode cs argTs bs xs rec_const params intr_ts mono
+ fp_def rec_preds_defs ctxt =
+ let
+ val _ = clean_message quiet_mode " Proving the induction rule ...";
+ val thy = ProofContext.theory_of ctxt;
+
+ (* predicates for induction rule *)
+
+ val (pnames, ctxt') = ctxt |>
+ Variable.add_fixes (map (fst o dest_Free) params) |> snd |>
+ Variable.variant_fixes (mk_names "P" (length cs));
+ val preds = map Free (pnames ~~
+ map (fn c => List.drop (binder_types (fastype_of c), length params) --->
+ HOLogic.boolT) cs);
+
+ (* transform an introduction rule into a premise for induction rule *)
+
+ fun mk_ind_prem r =
+ let
+ fun subst s = (case dest_predicate cs params s of
+ SOME (_, i, ys, (_, Ts)) =>
+ let
+ val k = length Ts;
+ val bs = map Bound (k - 1 downto 0);
+ val P = list_comb (List.nth (preds, i),
+ map (incr_boundvars k) ys @ bs);
+ val Q = list_abs (mk_names "x" k ~~ Ts,
+ HOLogic.mk_binop inductive_conj_name
+ (list_comb (incr_boundvars k s, bs), P))
+ in (Q, case Ts of [] => SOME (s, P) | _ => NONE) end
+ | NONE => (case s of
+ (t $ u) => (fst (subst t) $ fst (subst u), NONE)
+ | (Abs (a, T, t)) => (Abs (a, T, fst (subst t)), NONE)
+ | _ => (s, NONE)));
+
+ fun mk_prem (s, prems) = (case subst s of
+ (_, SOME (t, u)) => t :: u :: prems
+ | (t, _) => t :: prems);
+
+ val SOME (_, i, ys, _) = dest_predicate cs params
+ (HOLogic.dest_Trueprop (Logic.strip_assums_concl r))
+
+ in list_all_free (Logic.strip_params r,
+ Logic.list_implies (map HOLogic.mk_Trueprop (List.foldr mk_prem
+ [] (map HOLogic.dest_Trueprop (Logic.strip_assums_hyp r))),
+ HOLogic.mk_Trueprop (list_comb (List.nth (preds, i), ys))))
+ end;
+
+ val ind_prems = map mk_ind_prem intr_ts;
+
+
+ (* make conclusions for induction rules *)
+
+ val Tss = map (binder_types o fastype_of) preds;
+ val (xnames, ctxt'') =
+ Variable.variant_fixes (mk_names "x" (length (flat Tss))) ctxt';
+ val mutual_ind_concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+ (map (fn (((xnames, Ts), c), P) =>
+ let val frees = map Free (xnames ~~ Ts)
+ in HOLogic.mk_imp
+ (list_comb (c, params @ frees), list_comb (P, frees))
+ end) (unflat Tss xnames ~~ Tss ~~ cs ~~ preds)));
+
+
+ (* make predicate for instantiation of abstract induction rule *)
+
+ val ind_pred = fold_rev lambda (bs @ xs) (foldr1 HOLogic.mk_conj
+ (map_index (fn (i, P) => List.foldr HOLogic.mk_imp
+ (list_comb (P, make_args' argTs xs (binder_types (fastype_of P))))
+ (make_bool_args HOLogic.mk_not I bs i)) preds));
+
+ val ind_concl = HOLogic.mk_Trueprop
+ (HOLogic.mk_binrel "HOL.ord_class.less_eq" (rec_const, ind_pred));
+
+ val raw_fp_induct = (mono RS (fp_def RS def_lfp_induct));
+
+ val induct = SkipProof.prove ctxt'' [] ind_prems ind_concl
+ (fn {prems, ...} => EVERY
+ [rewrite_goals_tac [inductive_conj_def],
+ DETERM (rtac raw_fp_induct 1),
+ REPEAT (resolve_tac [le_funI, le_boolI] 1),
+ rewrite_goals_tac (inf_fun_eq :: inf_bool_eq :: simp_thms'),
+ (*This disjE separates out the introduction rules*)
+ REPEAT (FIRSTGOAL (eresolve_tac [disjE, exE, FalseE])),
+ (*Now break down the individual cases. No disjE here in case
+ some premise involves disjunction.*)
+ REPEAT (FIRSTGOAL (etac conjE ORELSE' bound_hyp_subst_tac)),
+ REPEAT (FIRSTGOAL
+ (resolve_tac [conjI, impI] ORELSE' (etac notE THEN' atac))),
+ EVERY (map (fn prem => DEPTH_SOLVE_1 (ares_tac [rewrite_rule
+ (inductive_conj_def :: rec_preds_defs @ simp_thms') prem,
+ conjI, refl] 1)) prems)]);
+
+ val lemma = SkipProof.prove ctxt'' [] []
+ (Logic.mk_implies (ind_concl, mutual_ind_concl)) (fn _ => EVERY
+ [rewrite_goals_tac rec_preds_defs,
+ REPEAT (EVERY
+ [REPEAT (resolve_tac [conjI, impI] 1),
+ REPEAT (eresolve_tac [le_funE, le_boolE] 1),
+ atac 1,
+ rewrite_goals_tac simp_thms',
+ atac 1])])
+
+ in singleton (ProofContext.export ctxt'' ctxt) (induct RS lemma) end;
+
+
+
+(** specification of (co)inductive predicates **)
+
+fun mk_ind_def quiet_mode skip_mono fork_mono alt_name coind cs intr_ts monos params cnames_syn ctxt =
+ let
+ val fp_name = if coind then @{const_name Inductive.gfp} else @{const_name Inductive.lfp};
+
+ val argTs = fold (fn c => fn Ts => Ts @
+ (List.drop (binder_types (fastype_of c), length params) \\ Ts)) cs [];
+ val k = log 2 1 (length cs);
+ val predT = replicate k HOLogic.boolT ---> argTs ---> HOLogic.boolT;
+ val p :: xs = map Free (Variable.variant_frees ctxt intr_ts
+ (("p", predT) :: (mk_names "x" (length argTs) ~~ argTs)));
+ val bs = map Free (Variable.variant_frees ctxt (p :: xs @ intr_ts)
+ (map (rpair HOLogic.boolT) (mk_names "b" k)));
+
+ fun subst t = (case dest_predicate cs params t of
+ SOME (_, i, ts, (Ts, Us)) =>
+ let
+ val l = length Us;
+ val zs = map Bound (l - 1 downto 0)
+ in
+ list_abs (map (pair "z") Us, list_comb (p,
+ make_bool_args' bs i @ make_args argTs
+ ((map (incr_boundvars l) ts ~~ Ts) @ (zs ~~ Us))))
+ end
+ | NONE => (case t of
+ t1 $ t2 => subst t1 $ subst t2
+ | Abs (x, T, u) => Abs (x, T, subst u)
+ | _ => t));
+
+ (* transform an introduction rule into a conjunction *)
+ (* [| p_i t; ... |] ==> p_j u *)
+ (* is transformed into *)
+ (* b_j & x_j = u & p b_j t & ... *)
+
+ fun transform_rule r =
+ let
+ val SOME (_, i, ts, (Ts, _)) = dest_predicate cs params
+ (HOLogic.dest_Trueprop (Logic.strip_assums_concl r));
+ val ps = make_bool_args HOLogic.mk_not I bs i @
+ map HOLogic.mk_eq (make_args' argTs xs Ts ~~ ts) @
+ map (subst o HOLogic.dest_Trueprop)
+ (Logic.strip_assums_hyp r)
+ in List.foldr (fn ((x, T), P) => HOLogic.exists_const T $ (Abs (x, T, P)))
+ (if null ps then HOLogic.true_const else foldr1 HOLogic.mk_conj ps)
+ (Logic.strip_params r)
+ end
+
+ (* make a disjunction of all introduction rules *)
+
+ val fp_fun = fold_rev lambda (p :: bs @ xs)
+ (if null intr_ts then HOLogic.false_const
+ else foldr1 HOLogic.mk_disj (map transform_rule intr_ts));
+
+ (* add definiton of recursive predicates to theory *)
+
+ val rec_name =
+ if Binding.is_empty alt_name then
+ Binding.name (space_implode "_" (map (Binding.name_of o fst) cnames_syn))
+ else alt_name;
+
+ val ((rec_const, (_, fp_def)), ctxt') = ctxt |>
+ LocalTheory.define Thm.internalK
+ ((rec_name, case cnames_syn of [(_, syn)] => syn | _ => NoSyn),
+ (Attrib.empty_binding, fold_rev lambda params
+ (Const (fp_name, (predT --> predT) --> predT) $ fp_fun)));
+ val fp_def' = Simplifier.rewrite (HOL_basic_ss addsimps [fp_def])
+ (cterm_of (ProofContext.theory_of ctxt') (list_comb (rec_const, params)));
+ val specs = if length cs < 2 then [] else
+ map_index (fn (i, (name_mx, c)) =>
+ let
+ val Ts = List.drop (binder_types (fastype_of c), length params);
+ val xs = map Free (Variable.variant_frees ctxt intr_ts
+ (mk_names "x" (length Ts) ~~ Ts))
+ in
+ (name_mx, (Attrib.empty_binding, fold_rev lambda (params @ xs)
+ (list_comb (rec_const, params @ make_bool_args' bs i @
+ make_args argTs (xs ~~ Ts)))))
+ end) (cnames_syn ~~ cs);
+ val (consts_defs, ctxt'') = fold_map (LocalTheory.define Thm.internalK) specs ctxt';
+ val preds = (case cs of [_] => [rec_const] | _ => map #1 consts_defs);
+
+ val mono = prove_mono quiet_mode skip_mono fork_mono predT fp_fun monos ctxt'';
+ val ((_, [mono']), ctxt''') =
+ LocalTheory.note Thm.internalK (Attrib.empty_binding, [mono]) ctxt'';
+
+ in (ctxt''', rec_name, mono', fp_def', map (#2 o #2) consts_defs,
+ list_comb (rec_const, params), preds, argTs, bs, xs)
+ end;
+
+fun declare_rules kind rec_binding coind no_ind cnames intrs intr_bindings intr_atts
+ elims raw_induct ctxt =
+ let
+ val rec_name = Binding.name_of rec_binding;
+ val rec_qualified = Binding.qualify false rec_name;
+ val intr_names = map Binding.name_of intr_bindings;
+ val ind_case_names = RuleCases.case_names intr_names;
+ val induct =
+ if coind then
+ (raw_induct, [RuleCases.case_names [rec_name],
+ RuleCases.case_conclusion (rec_name, intr_names),
+ RuleCases.consumes 1, Induct.coinduct_pred (hd cnames)])
+ else if no_ind orelse length cnames > 1 then
+ (raw_induct, [ind_case_names, RuleCases.consumes 0])
+ else (raw_induct RSN (2, rev_mp), [ind_case_names, RuleCases.consumes 1]);
+
+ val (intrs', ctxt1) =
+ ctxt |>
+ LocalTheory.notes kind
+ (map rec_qualified intr_bindings ~~ intr_atts ~~ map (fn th => [([th],
+ [Attrib.internal (K (ContextRules.intro_query NONE)),
+ Attrib.internal (K Nitpick_Ind_Intro_Thms.add)])]) intrs) |>>
+ map (hd o snd);
+ val (((_, elims'), (_, [induct'])), ctxt2) =
+ ctxt1 |>
+ LocalTheory.note kind ((rec_qualified (Binding.name "intros"), []), intrs') ||>>
+ fold_map (fn (name, (elim, cases)) =>
+ LocalTheory.note kind ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "cases"),
+ [Attrib.internal (K (RuleCases.case_names cases)),
+ Attrib.internal (K (RuleCases.consumes 1)),
+ Attrib.internal (K (Induct.cases_pred name)),
+ Attrib.internal (K (ContextRules.elim_query NONE))]), [elim]) #>
+ apfst (hd o snd)) (if null elims then [] else cnames ~~ elims) ||>>
+ LocalTheory.note kind
+ ((rec_qualified (Binding.name (coind_prefix coind ^ "induct")),
+ map (Attrib.internal o K) (#2 induct)), [rulify (#1 induct)]);
+
+ val ctxt3 = if no_ind orelse coind then ctxt2 else
+ let val inducts = cnames ~~ ProjectRule.projects ctxt2 (1 upto length cnames) induct'
+ in
+ ctxt2 |>
+ LocalTheory.notes kind [((rec_qualified (Binding.name "inducts"), []),
+ inducts |> map (fn (name, th) => ([th],
+ [Attrib.internal (K ind_case_names),
+ Attrib.internal (K (RuleCases.consumes 1)),
+ Attrib.internal (K (Induct.induct_pred name))])))] |> snd
+ end
+ in (intrs', elims', induct', ctxt3) end;
+
+type inductive_flags =
+ {quiet_mode: bool, verbose: bool, kind: string, alt_name: binding,
+ coind: bool, no_elim: bool, no_ind: bool, skip_mono: bool, fork_mono: bool}
+
+type add_ind_def =
+ inductive_flags ->
+ term list -> (Attrib.binding * term) list -> thm list ->
+ term list -> (binding * mixfix) list ->
+ local_theory -> inductive_result * local_theory
+
+fun add_ind_def {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono}
+ cs intros monos params cnames_syn ctxt =
+ let
+ val _ = null cnames_syn andalso error "No inductive predicates given";
+ val names = map (Binding.name_of o fst) cnames_syn;
+ val _ = message (quiet_mode andalso not verbose)
+ ("Proofs for " ^ coind_prefix coind ^ "inductive predicate(s) " ^ commas_quote names);
+
+ val cnames = map (LocalTheory.full_name ctxt o #1) cnames_syn; (* FIXME *)
+ val ((intr_names, intr_atts), intr_ts) =
+ apfst split_list (split_list (map (check_rule ctxt cs params) intros));
+
+ val (ctxt1, rec_name, mono, fp_def, rec_preds_defs, rec_const, preds,
+ argTs, bs, xs) = mk_ind_def quiet_mode skip_mono fork_mono alt_name coind cs intr_ts
+ monos params cnames_syn ctxt;
+
+ val (intrs, unfold) = prove_intrs quiet_mode coind mono fp_def (length bs + length xs)
+ params intr_ts rec_preds_defs ctxt1;
+ val elims = if no_elim then [] else
+ prove_elims quiet_mode cs params intr_ts (map Binding.name_of intr_names)
+ unfold rec_preds_defs ctxt1;
+ val raw_induct = zero_var_indexes
+ (if no_ind then Drule.asm_rl else
+ if coind then
+ singleton (ProofContext.export
+ (snd (Variable.add_fixes (map (fst o dest_Free) params) ctxt1)) ctxt1)
+ (rotate_prems ~1 (ObjectLogic.rulify
+ (fold_rule rec_preds_defs
+ (rewrite_rule [le_fun_def, le_bool_def, sup_fun_eq, sup_bool_eq]
+ (mono RS (fp_def RS def_coinduct))))))
+ else
+ prove_indrule quiet_mode cs argTs bs xs rec_const params intr_ts mono fp_def
+ rec_preds_defs ctxt1);
+
+ val (intrs', elims', induct, ctxt2) = declare_rules kind rec_name coind no_ind
+ cnames intrs intr_names intr_atts elims raw_induct ctxt1;
+
+ val result =
+ {preds = preds,
+ intrs = intrs',
+ elims = elims',
+ raw_induct = rulify raw_induct,
+ induct = induct};
+
+ val ctxt3 = ctxt2
+ |> LocalTheory.declaration (fn phi =>
+ let val result' = morph_result phi result;
+ in put_inductives cnames (*global names!?*) ({names = cnames, coind = coind}, result') end);
+ in (result, ctxt3) end;
+
+
+(* external interfaces *)
+
+fun gen_add_inductive_i mk_def
+ (flags as {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono})
+ cnames_syn pnames spec monos lthy =
+ let
+ val thy = ProofContext.theory_of lthy;
+ val _ = Theory.requires thy "Inductive" (coind_prefix coind ^ "inductive definitions");
+
+
+ (* abbrevs *)
+
+ val (_, ctxt1) = Variable.add_fixes (map (Binding.name_of o fst o fst) cnames_syn) lthy;
+
+ fun get_abbrev ((name, atts), t) =
+ if can (Logic.strip_assums_concl #> Logic.dest_equals) t then
+ let
+ val _ = Binding.is_empty name andalso null atts orelse
+ error "Abbreviations may not have names or attributes";
+ val ((x, T), rhs) = LocalDefs.abs_def (snd (LocalDefs.cert_def ctxt1 t));
+ val var =
+ (case find_first (fn ((c, _), _) => Binding.name_of c = x) cnames_syn of
+ NONE => error ("Undeclared head of abbreviation " ^ quote x)
+ | SOME ((b, T'), mx) =>
+ if T <> T' then error ("Bad type specification for abbreviation " ^ quote x)
+ else (b, mx));
+ in SOME (var, rhs) end
+ else NONE;
+
+ val abbrevs = map_filter get_abbrev spec;
+ val bs = map (Binding.name_of o fst o fst) abbrevs;
+
+
+ (* predicates *)
+
+ val pre_intros = filter_out (is_some o get_abbrev) spec;
+ val cnames_syn' = filter_out (member (op =) bs o Binding.name_of o fst o fst) cnames_syn;
+ val cs = map (Free o apfst Binding.name_of o fst) cnames_syn';
+ val ps = map Free pnames;
+
+ val (_, ctxt2) = lthy |> Variable.add_fixes (map (Binding.name_of o fst o fst) cnames_syn');
+ val _ = map (fn abbr => LocalDefs.fixed_abbrev abbr ctxt2) abbrevs;
+ val ctxt3 = ctxt2 |> fold (snd oo LocalDefs.fixed_abbrev) abbrevs;
+ val expand = Assumption.export_term ctxt3 lthy #> ProofContext.cert_term lthy;
+
+ fun close_rule r = list_all_free (rev (fold_aterms
+ (fn t as Free (v as (s, _)) =>
+ if Variable.is_fixed ctxt1 s orelse
+ member (op =) ps t then I else insert (op =) v
+ | _ => I) r []), r);
+
+ val intros = map (apsnd (Syntax.check_term lthy #> close_rule #> expand)) pre_intros;
+ val preds = map (fn ((c, _), mx) => (c, mx)) cnames_syn';
+ in
+ lthy
+ |> mk_def flags cs intros monos ps preds
+ ||> fold (snd oo LocalTheory.abbrev Syntax.mode_default) abbrevs
+ end;
+
+fun gen_add_inductive mk_def verbose coind cnames_syn pnames_syn intro_srcs raw_monos int lthy =
+ let
+ val ((vars, intrs), _) = lthy
+ |> ProofContext.set_mode ProofContext.mode_abbrev
+ |> Specification.read_spec (cnames_syn @ pnames_syn) intro_srcs;
+ val (cs, ps) = chop (length cnames_syn) vars;
+ val monos = Attrib.eval_thms lthy raw_monos;
+ val flags = {quiet_mode = false, verbose = verbose, kind = Thm.generatedK,
+ alt_name = Binding.empty, coind = coind, no_elim = false, no_ind = false,
+ skip_mono = false, fork_mono = not int};
+ in
+ lthy
+ |> LocalTheory.set_group (serial_string ())
+ |> gen_add_inductive_i mk_def flags cs (map (apfst Binding.name_of o fst) ps) intrs monos
+ end;
+
+val add_inductive_i = gen_add_inductive_i add_ind_def;
+val add_inductive = gen_add_inductive add_ind_def;
+
+fun add_inductive_global group flags cnames_syn pnames pre_intros monos thy =
+ let
+ val name = Sign.full_name thy (fst (fst (hd cnames_syn)));
+ val ctxt' = thy
+ |> TheoryTarget.init NONE
+ |> LocalTheory.set_group group
+ |> add_inductive_i flags cnames_syn pnames pre_intros monos |> snd
+ |> LocalTheory.exit;
+ val info = #2 (the_inductive ctxt' name);
+ in (info, ProofContext.theory_of ctxt') end;
+
+
+(* read off arities of inductive predicates from raw induction rule *)
+fun arities_of induct =
+ map (fn (_ $ t $ u) =>
+ (fst (dest_Const (head_of t)), length (snd (strip_comb u))))
+ (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
+
+(* read off parameters of inductive predicate from raw induction rule *)
+fun params_of induct =
+ let
+ val (_ $ t $ u :: _) =
+ HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct));
+ val (_, ts) = strip_comb t;
+ val (_, us) = strip_comb u
+ in
+ List.take (ts, length ts - length us)
+ end;
+
+val pname_of_intr =
+ concl_of #> HOLogic.dest_Trueprop #> head_of #> dest_Const #> fst;
+
+(* partition introduction rules according to predicate name *)
+fun gen_partition_rules f induct intros =
+ fold_rev (fn r => AList.map_entry op = (pname_of_intr (f r)) (cons r)) intros
+ (map (rpair [] o fst) (arities_of induct));
+
+val partition_rules = gen_partition_rules I;
+fun partition_rules' induct = gen_partition_rules fst induct;
+
+fun unpartition_rules intros xs =
+ fold_map (fn r => AList.map_entry_yield op = (pname_of_intr r)
+ (fn x :: xs => (x, xs)) #>> the) intros xs |> fst;
+
+(* infer order of variables in intro rules from order of quantifiers in elim rule *)
+fun infer_intro_vars elim arity intros =
+ let
+ val thy = theory_of_thm elim;
+ val _ :: cases = prems_of elim;
+ val used = map (fst o fst) (Term.add_vars (prop_of elim) []);
+ fun mtch (t, u) =
+ let
+ val params = Logic.strip_params t;
+ val vars = map (Var o apfst (rpair 0))
+ (Name.variant_list used (map fst params) ~~ map snd params);
+ val ts = map (curry subst_bounds (rev vars))
+ (List.drop (Logic.strip_assums_hyp t, arity));
+ val us = Logic.strip_imp_prems u;
+ val tab = fold (Pattern.first_order_match thy) (ts ~~ us)
+ (Vartab.empty, Vartab.empty);
+ in
+ map (Envir.subst_vars tab) vars
+ end
+ in
+ map (mtch o apsnd prop_of) (cases ~~ intros)
+ end;
+
+
+
+(** package setup **)
+
+(* setup theory *)
+
+val setup =
+ ind_cases_setup #>
+ Attrib.setup @{binding mono} (Attrib.add_del mono_add mono_del)
+ "declaration of monotonicity rule";
+
+
+(* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val _ = OuterKeyword.keyword "monos";
+
+fun gen_ind_decl mk_def coind =
+ P.fixes -- P.for_fixes --
+ Scan.optional SpecParse.where_alt_specs [] --
+ Scan.optional (P.$$$ "monos" |-- P.!!! SpecParse.xthms1) []
+ >> (fn (((preds, params), specs), monos) =>
+ (snd oo gen_add_inductive mk_def true coind preds params specs monos));
+
+val ind_decl = gen_ind_decl add_ind_def;
+
+val _ = OuterSyntax.local_theory' "inductive" "define inductive predicates" K.thy_decl (ind_decl false);
+val _ = OuterSyntax.local_theory' "coinductive" "define coinductive predicates" K.thy_decl (ind_decl true);
+
+val _ =
+ OuterSyntax.local_theory "inductive_cases"
+ "create simplified instances of elimination rules (improper)" K.thy_script
+ (P.and_list1 SpecParse.specs >> (snd oo inductive_cases));
+
+end;
+
+end;
--- a/src/HOL/Tools/inductive_codegen.ML Thu Jun 18 18:31:14 2009 -0700
+++ b/src/HOL/Tools/inductive_codegen.ML Fri Jun 19 17:23:21 2009 +0200
@@ -62,9 +62,9 @@
val nparms = (case optnparms of
SOME k => k
| NONE => (case rules of
- [] => (case try (InductivePackage.the_inductive (ProofContext.init thy)) s of
+ [] => (case try (Inductive.the_inductive (ProofContext.init thy)) s of
SOME (_, {raw_induct, ...}) =>
- length (InductivePackage.params_of raw_induct)
+ length (Inductive.params_of raw_induct)
| NONE => 0)
| xs => snd (snd (snd (split_last xs)))))
in CodegenData.put
@@ -81,11 +81,11 @@
fun get_clauses thy s =
let val {intros, graph, ...} = CodegenData.get thy
in case Symtab.lookup intros s of
- NONE => (case try (InductivePackage.the_inductive (ProofContext.init thy)) s of
+ NONE => (case try (Inductive.the_inductive (ProofContext.init thy)) s of
NONE => NONE
| SOME ({names, ...}, {intrs, raw_induct, ...}) =>
SOME (names, Codegen.thyname_of_const thy s,
- length (InductivePackage.params_of raw_induct),
+ length (Inductive.params_of raw_induct),
preprocess thy intrs))
| SOME _ =>
let
@@ -103,7 +103,7 @@
let
val cnstrs = List.concat (List.concat (map
(map (fn (_, (_, _, cs)) => map (apsnd length) cs) o #descr o snd)
- (Symtab.dest (DatatypePackage.get_datatypes thy))));
+ (Symtab.dest (Datatype.get_datatypes thy))));
fun check t = (case strip_comb t of
(Var _, []) => true
| (Const (s, _), ts) => (case AList.lookup (op =) cnstrs s of
--- a/src/HOL/Tools/inductive_package.ML Thu Jun 18 18:31:14 2009 -0700
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,968 +0,0 @@
-(* Title: HOL/Tools/inductive_package.ML
- Author: Lawrence C Paulson, Cambridge University Computer Laboratory
- Author: Stefan Berghofer and Markus Wenzel, TU Muenchen
-
-(Co)Inductive Definition module for HOL.
-
-Features:
- * least or greatest fixedpoints
- * mutually recursive definitions
- * definitions involving arbitrary monotone operators
- * automatically proves introduction and elimination rules
-
- Introduction rules have the form
- [| M Pj ti, ..., Q x, ... |] ==> Pk t
- where M is some monotone operator (usually the identity)
- Q x is any side condition on the free variables
- ti, t are any terms
- Pj, Pk are two of the predicates being defined in mutual recursion
-*)
-
-signature BASIC_INDUCTIVE_PACKAGE =
-sig
- type inductive_result
- val morph_result: morphism -> inductive_result -> inductive_result
- type inductive_info
- val the_inductive: Proof.context -> string -> inductive_info
- val print_inductives: Proof.context -> unit
- val mono_add: attribute
- val mono_del: attribute
- val get_monos: Proof.context -> thm list
- val mk_cases: Proof.context -> term -> thm
- val inductive_forall_name: string
- val inductive_forall_def: thm
- val rulify: thm -> thm
- val inductive_cases: (Attrib.binding * string list) list -> local_theory ->
- thm list list * local_theory
- val inductive_cases_i: (Attrib.binding * term list) list -> local_theory ->
- thm list list * local_theory
- type inductive_flags
- val add_inductive_i:
- inductive_flags -> ((binding * typ) * mixfix) list ->
- (st