src/HOL/Tools/typedef.ML
changeset 35741 4f3660a3e5af
parent 35682 5e6811f4294b
child 35766 eafaa9990712
equal deleted inserted replaced
35740:d3726291f252 35741:4f3660a3e5af
     1 (*  Title:      HOL/Tools/typedef.ML
     1 (*  Title:      HOL/Tools/typedef.ML
     2     Author:     Markus Wenzel and Stefan Berghofer, TU Muenchen
     2     Author:     Markus Wenzel and Stefan Berghofer, TU Muenchen
     3 
     3 
     4 Gordon/HOL-style type definitions: create a new syntactic type
     4 Gordon/HOL-style type definitions: create a new syntactic type
     5 represented by a non-empty subset.
     5 represented by a non-empty set.
     6 *)
     6 *)
     7 
     7 
     8 signature TYPEDEF =
     8 signature TYPEDEF =
     9 sig
     9 sig
    10   type info =
    10   type info =
    11    {rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string, inhabited: thm,
    11    {rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string, inhabited: thm,
    12     type_definition: thm, set_def: thm option, Rep: thm, Rep_inverse: thm,
    12     type_definition: thm, set_def: thm option, Rep: thm, Rep_inverse: thm,
    13     Abs_inverse: thm, Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm,
    13     Abs_inverse: thm, Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm,
    14     Rep_induct: thm, Abs_induct: thm}
    14     Rep_induct: thm, Abs_induct: thm}
       
    15   val transform_info: morphism -> info -> info
       
    16   val get_info: Proof.context -> string -> info list
       
    17   val get_info_global: theory -> string -> info list
       
    18   val interpretation: (string -> theory -> theory) -> theory -> theory
       
    19   val setup: theory -> theory
    15   val add_typedef: bool -> binding option -> binding * string list * mixfix ->
    20   val add_typedef: bool -> binding option -> binding * string list * mixfix ->
       
    21     term -> (binding * binding) option -> tactic -> local_theory -> (string * info) * local_theory
       
    22   val add_typedef_global: bool -> binding option -> binding * string list * mixfix ->
    16     term -> (binding * binding) option -> tactic -> theory -> (string * info) * theory
    23     term -> (binding * binding) option -> tactic -> theory -> (string * info) * theory
    17   val typedef: (bool * binding) * (binding * string list * mixfix) * term *
    24   val typedef: (bool * binding) * (binding * string list * mixfix) * term *
    18     (binding * binding) option -> theory -> Proof.state
    25     (binding * binding) option -> local_theory -> Proof.state
    19   val typedef_cmd: (bool * binding) * (binding * string list * mixfix) * string *
    26   val typedef_cmd: (bool * binding) * (binding * string list * mixfix) * string *
    20     (binding * binding) option -> theory -> Proof.state
    27     (binding * binding) option -> local_theory -> Proof.state
    21   val get_info: theory -> string -> info option
       
    22   val the_info: theory -> string -> info
       
    23   val interpretation: (string -> theory -> theory) -> theory -> theory
       
    24   val setup: theory -> theory
       
    25 end;
    28 end;
    26 
    29 
    27 structure Typedef: TYPEDEF =
    30 structure Typedef: TYPEDEF =
    28 struct
    31 struct
    29 
    32 
    30 (** type definitions **)
    33 (** type definitions **)
    31 
    34 
    32 (* theory data *)
    35 (* theory data *)
    33 
    36 
    34 type info =
    37 type info =
    35  {rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string, inhabited: thm,
    38  {(*global part*)
    36   type_definition: thm, set_def: thm option, Rep: thm, Rep_inverse: thm,
    39   rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string,
       
    40   (*local part*)
       
    41   inhabited: thm, type_definition: thm, set_def: thm option, Rep: thm, Rep_inverse: thm,
    37   Abs_inverse: thm, Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm,
    42   Abs_inverse: thm, Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm,
    38   Rep_induct: thm, Abs_induct: thm};
    43   Rep_induct: thm, Abs_induct: thm};
    39 
    44 
    40 structure TypedefData = Theory_Data
    45 fun transform_info phi (info: info) =
       
    46   let
       
    47     val thm = Morphism.thm phi;
       
    48     val {rep_type, abs_type, Rep_name, Abs_name, inhabited, type_definition,
       
    49       set_def, Rep, Rep_inverse, Abs_inverse, Rep_inject, Abs_inject,
       
    50       Rep_cases, Abs_cases, Rep_induct, Abs_induct} = info;
       
    51   in
       
    52    {rep_type = rep_type, abs_type = abs_type, Rep_name = Rep_name, Abs_name = Abs_name,
       
    53     inhabited = thm inhabited, type_definition = thm type_definition,
       
    54     set_def = Option.map thm set_def, Rep = thm Rep, Rep_inverse = thm Rep_inverse,
       
    55     Abs_inverse = thm Abs_inverse, Rep_inject = thm Rep_inject, Abs_inject = thm Abs_inject,
       
    56     Rep_cases = thm Rep_cases, Abs_cases = thm Abs_cases, Rep_induct = thm Rep_induct,
       
    57     Abs_induct = thm Abs_induct}
       
    58   end;
       
    59 
       
    60 structure Data = Generic_Data
    41 (
    61 (
    42   type T = info Symtab.table;
    62   type T = info list Symtab.table;
    43   val empty = Symtab.empty;
    63   val empty = Symtab.empty;
    44   val extend = I;
    64   val extend = I;
    45   fun merge data = Symtab.merge (K true) data;
    65   fun merge data = Symtab.merge_list (K true) data;
    46 );
    66 );
    47 
    67 
    48 val get_info = Symtab.lookup o TypedefData.get;
    68 val get_info = Symtab.lookup_list o Data.get o Context.Proof;
    49 
    69 val get_info_global = Symtab.lookup_list o Data.get o Context.Theory;
    50 fun the_info thy name =
    70 
    51   (case get_info thy name of
    71 fun put_info name info = Data.map (Symtab.cons_list (name, info));
    52     SOME info => info
    72 
    53   | NONE => error ("Unknown typedef " ^ quote name));
    73 
    54 
    74 (* global interpretation *)
    55 fun put_info name info = TypedefData.map (Symtab.update (name, info));
       
    56 
       
    57 
       
    58 (* prepare_typedef *)
       
    59 
       
    60 fun declare_type_name a = Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS)));
       
    61 
    75 
    62 structure Typedef_Interpretation = Interpretation(type T = string val eq = op =);
    76 structure Typedef_Interpretation = Interpretation(type T = string val eq = op =);
    63 val interpretation = Typedef_Interpretation.interpretation;
    77 val interpretation = Typedef_Interpretation.interpretation;
    64 
    78 
    65 fun prepare_typedef prep_term def name (tname, vs, mx) raw_set opt_morphs thy =
    79 val setup = Typedef_Interpretation.init;
    66   let
    80 
    67     val _ = Theory.requires thy "Typedef" "typedefs";
    81 
    68     val ctxt = ProofContext.init thy;
    82 (* primitive typedef axiomatization -- for fresh typedecl *)
    69 
    83 
    70     val full = Sign.full_name thy;
    84 fun mk_inhabited A =
    71     val full_name = full name;
    85   let val T = HOLogic.dest_setT (Term.fastype_of A)
       
    86   in HOLogic.mk_Trueprop (HOLogic.exists_const T $ Abs ("x", T, HOLogic.mk_mem (Bound 0, A))) end;
       
    87 
       
    88 fun mk_typedef newT oldT RepC AbsC A =
       
    89   let
       
    90     val typedefC =
       
    91       Const (@{const_name type_definition},
       
    92         (newT --> oldT) --> (oldT --> newT) --> HOLogic.mk_setT oldT --> HOLogic.boolT);
       
    93   in Logic.mk_implies (mk_inhabited A, HOLogic.mk_Trueprop (typedefC $ RepC $ AbsC $ A)) end;
       
    94 
       
    95 fun primitive_typedef typedef_name newT oldT Rep_name Abs_name A thy =
       
    96   let
       
    97     (* errors *)
       
    98 
       
    99     fun show_names pairs = commas_quote (map fst pairs);
       
   100 
       
   101     val lhs_tfrees = Term.add_tfreesT newT [];
       
   102     val rhs_tfrees = Term.add_tfreesT oldT [];
       
   103     val _ =
       
   104       (case fold (remove (op =)) lhs_tfrees rhs_tfrees of [] => ()
       
   105       | extras => error ("Extra type variables in representing set: " ^ show_names extras));
       
   106 
       
   107     val _ =
       
   108       (case Term.add_frees A [] of [] => []
       
   109       | xs => error ("Illegal variables in representing set: " ^ show_names xs));
       
   110 
       
   111 
       
   112     (* axiomatization *)
       
   113 
       
   114     val ((RepC, AbsC), consts_thy) = thy
       
   115       |> Sign.declare_const ((Rep_name, newT --> oldT), NoSyn)
       
   116       ||>> Sign.declare_const ((Abs_name, oldT --> newT), NoSyn);
       
   117 
       
   118     val typedef_deps = Term.add_consts A [];
       
   119 
       
   120     val (axiom, axiom_thy) = consts_thy
       
   121       |> Thm.add_axiom (typedef_name, mk_typedef newT oldT RepC AbsC A)
       
   122       ||> Theory.add_deps "" (dest_Const RepC) typedef_deps
       
   123       ||> Theory.add_deps "" (dest_Const AbsC) typedef_deps;
       
   124 
       
   125   in ((RepC, AbsC, axiom), axiom_thy) end;
       
   126 
       
   127 
       
   128 (* prepare_typedef *)
       
   129 
       
   130 fun declare_type_name a = Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS)));
       
   131 
       
   132 fun prepare_typedef prep_term def_set name (tname, vs, mx) raw_set opt_morphs lthy =
       
   133   let
       
   134     val full_name = Local_Theory.full_name lthy name;
    72     val bname = Binding.name_of name;
   135     val bname = Binding.name_of name;
    73 
   136 
    74     (*rhs*)
   137 
    75     val set = prep_term (ctxt |> fold declare_type_name vs) raw_set;
   138     (* rhs *)
       
   139 
       
   140     val set = prep_term (lthy |> fold declare_type_name vs) raw_set;
    76     val setT = Term.fastype_of set;
   141     val setT = Term.fastype_of set;
       
   142     val oldT = HOLogic.dest_setT setT handle TYPE _ =>
       
   143       error ("Not a set type: " ^ quote (Syntax.string_of_typ lthy setT));
       
   144 
       
   145     val goal = mk_inhabited set;
       
   146     val goal_pat = mk_inhabited (Var (the_default (bname, 0) (Syntax.read_variable bname), setT));
       
   147 
       
   148 
       
   149     (* lhs *)
       
   150 
       
   151     val (newT, typedecl_lthy) = lthy
       
   152       |> Typedecl.typedecl_wrt [set] (tname, vs, mx)
       
   153       ||> Variable.declare_term set;
       
   154 
       
   155     val Type (full_tname, type_args) = newT;
       
   156     val lhs_tfrees = map Term.dest_TFree type_args;
       
   157 
       
   158 
       
   159     (* set definition *)
       
   160 
       
   161     (* FIXME let Local_Theory.define handle hidden polymorphism (!??!) *)
       
   162 
    77     val rhs_tfrees = Term.add_tfrees set [];
   163     val rhs_tfrees = Term.add_tfrees set [];
    78     val rhs_tfreesT = Term.add_tfreesT setT [];
   164     val rhs_tfreesT = Term.add_tfreesT setT [];
    79     val oldT = HOLogic.dest_setT setT handle TYPE _ =>
   165 
    80       error ("Not a set type: " ^ quote (Syntax.string_of_typ ctxt setT));
   166     val set_argsT = lhs_tfrees
    81 
       
    82     (*lhs*)
       
    83     val defS = Sign.defaultS thy;
       
    84     val lhs_tfrees = map (fn v => (v, the_default defS (AList.lookup (op =) rhs_tfrees v))) vs;
       
    85     val args_setT = lhs_tfrees
       
    86       |> filter (member (op =) rhs_tfrees andf (not o member (op =) rhs_tfreesT))
   167       |> filter (member (op =) rhs_tfrees andf (not o member (op =) rhs_tfreesT))
    87       |> map TFree;
   168       |> map TFree;
    88 
   169     val set_args = map Logic.mk_type set_argsT;
    89     val full_tname = full tname;
   170 
    90     val newT = Type (full_tname, map TFree lhs_tfrees);
   171     val ((set', set_def), set_lthy) =
       
   172       if def_set then
       
   173         typedecl_lthy
       
   174         |> Local_Theory.define
       
   175           ((name, NoSyn), ((Thm.def_binding name, []), fold_rev lambda set_args set))
       
   176         |>> (fn (s, (_, set_def)) => (Term.list_comb (s, set_args), SOME set_def))
       
   177       else ((set, NONE), typedecl_lthy);
       
   178 
       
   179 
       
   180     (* axiomatization *)
    91 
   181 
    92     val (Rep_name, Abs_name) =
   182     val (Rep_name, Abs_name) =
    93       (case opt_morphs of
   183       (case opt_morphs of
    94         NONE => (Binding.prefix_name "Rep_" name, Binding.prefix_name "Abs_" name)
   184         NONE => (Binding.prefix_name "Rep_" name, Binding.prefix_name "Abs_" name)
    95       | SOME morphs => morphs);
   185       | SOME morphs => morphs);
    96     val setT' = map Term.itselfT args_setT ---> setT;
   186 
    97     val setC = Term.list_comb (Const (full_name, setT'), map Logic.mk_type args_setT);
       
    98     val RepC = Const (full Rep_name, newT --> oldT);
       
    99     val AbsC = Const (full Abs_name, oldT --> newT);
       
   100 
       
   101     (*inhabitance*)
       
   102     fun mk_inhabited A =
       
   103       HOLogic.mk_Trueprop (HOLogic.mk_exists ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), A)));
       
   104     val set' = if def then setC else set;
       
   105     val goal' = mk_inhabited set';
       
   106     val goal = mk_inhabited set;
       
   107     val goal_pat = mk_inhabited (Var (the_default (bname, 0) (Syntax.read_variable bname), setT));
       
   108 
       
   109     (*axiomatization*)
       
   110     val typedef_name = Binding.prefix_name "type_definition_" name;
   187     val typedef_name = Binding.prefix_name "type_definition_" name;
   111     val typedefC =
   188 
   112       Const (@{const_name type_definition},
   189     val ((RepC, AbsC, typedef), typedef_lthy) =
   113         (newT --> oldT) --> (oldT --> newT) --> setT --> HOLogic.boolT);
   190       let
   114     val typedef_prop = Logic.mk_implies (goal', HOLogic.mk_Trueprop (typedefC $ RepC $ AbsC $ set'));
   191         val thy = ProofContext.theory_of set_lthy;
   115     val typedef_deps = Term.add_consts set' [];
   192         val cert = Thm.cterm_of thy;
   116 
   193         val (defs, A) =
   117     (*set definition*)
   194           Local_Defs.export_cterm set_lthy (ProofContext.init thy) (cert set') ||> Thm.term_of;
   118     fun add_def theory =
   195 
   119       if def then
   196         val ((RepC, AbsC, axiom), axiom_lthy) = set_lthy |>
   120         theory
   197           Local_Theory.theory_result (primitive_typedef typedef_name newT oldT Rep_name Abs_name A);
   121         |> Sign.add_consts_i [(name, setT', NoSyn)]
   198 
   122         |> PureThy.add_defs false [((Thm.def_binding name, Logic.mk_equals (setC, set)), [])]
   199         val cert = Thm.cterm_of (ProofContext.theory_of axiom_lthy);
   123         |-> (fn [th] => pair (SOME th))
   200         val typedef =
   124       else (NONE, theory);
   201           Local_Defs.contract axiom_lthy defs (cert (mk_typedef newT oldT RepC AbsC set')) axiom;
   125     fun contract_def NONE th = th
   202       in ((RepC, AbsC, typedef), axiom_lthy) end;
   126       | contract_def (SOME def_eq) th =
   203 
   127           let
   204     val alias_lthy = typedef_lthy
   128             val cert = Thm.cterm_of (Thm.theory_of_thm def_eq);
   205       |> Local_Theory.const_alias Rep_name (#1 (Term.dest_Const RepC))
   129             val goal_eq = MetaSimplifier.rewrite true [def_eq] (cert goal');
   206       |> Local_Theory.const_alias Abs_name (#1 (Term.dest_Const AbsC));
   130           in Drule.export_without_context (Drule.equal_elim_rule2 OF [goal_eq, th]) end;
   207 
   131 
   208 
   132     fun typedef_result inhabited =
   209     (* result *)
   133       Typedecl.typedecl_global (tname, vs, mx)
   210 
   134       #> snd
   211     fun note_qualify ((b, atts), th) =
   135       #> Sign.add_consts_i
   212       Local_Theory.note ((Binding.qualify false bname b, map (Attrib.internal o K) atts), [th])
   136         [(Rep_name, newT --> oldT, NoSyn),
   213       #>> (fn (_, [th']) => th');
   137          (Abs_name, oldT --> newT, NoSyn)]
   214 
   138       #> add_def
   215     fun typedef_result inhabited lthy1 =
   139       #-> (fn set_def =>
   216       let
   140         PureThy.add_axioms [((typedef_name, typedef_prop),
   217         val cert = Thm.cterm_of (ProofContext.theory_of lthy1);
   141           [Thm.rule_attribute (K (fn cond_axm => contract_def set_def inhabited RS cond_axm))])]
   218         val inhabited' =
   142         ##>> pair set_def)
   219           Local_Defs.contract lthy1 (the_list set_def) (cert (mk_inhabited set')) inhabited;
   143       ##> Theory.add_deps "" (dest_Const RepC) typedef_deps
   220         val typedef' = inhabited' RS typedef;
   144       ##> Theory.add_deps "" (dest_Const AbsC) typedef_deps
   221         fun make th = Goal.norm_result (typedef' RS th);
   145       #-> (fn ([type_definition], set_def) => fn thy1 =>
   222         val (((((((((((_, [type_definition]), Rep), Rep_inverse), Abs_inverse), Rep_inject),
   146         let
   223             Abs_inject), Rep_cases), Abs_cases), Rep_induct), Abs_induct), lthy2) = lthy1
   147           fun make th = Drule.export_without_context (th OF [type_definition]);
   224           |> Local_Theory.note ((typedef_name, []), [typedef'])
   148           val ([Rep, Rep_inverse, Abs_inverse, Rep_inject, Abs_inject,
   225           ||>> note_qualify ((Rep_name, []), make @{thm type_definition.Rep})
   149               Rep_cases, Abs_cases, Rep_induct, Abs_induct], thy2) =
   226           ||>> note_qualify ((Binding.suffix_name "_inverse" Rep_name, []),
   150             thy1
   227               make @{thm type_definition.Rep_inverse})
   151             |> Sign.add_path (Binding.name_of name)
   228           ||>> note_qualify ((Binding.suffix_name "_inverse" Abs_name, []),
   152             |> PureThy.add_thms
   229               make @{thm type_definition.Abs_inverse})
   153               [((Rep_name, make @{thm type_definition.Rep}), []),
   230           ||>> note_qualify ((Binding.suffix_name "_inject" Rep_name, []),
   154                 ((Binding.suffix_name "_inverse" Rep_name, make @{thm type_definition.Rep_inverse}), []),
   231               make @{thm type_definition.Rep_inject})
   155                 ((Binding.suffix_name "_inverse" Abs_name, make @{thm type_definition.Abs_inverse}), []),
   232           ||>> note_qualify ((Binding.suffix_name "_inject" Abs_name, []),
   156                 ((Binding.suffix_name "_inject" Rep_name, make @{thm type_definition.Rep_inject}), []),
   233               make @{thm type_definition.Abs_inject})
   157                 ((Binding.suffix_name "_inject" Abs_name, make @{thm type_definition.Abs_inject}), []),
   234           ||>> note_qualify ((Binding.suffix_name "_cases" Rep_name,
   158                 ((Binding.suffix_name "_cases" Rep_name, make @{thm type_definition.Rep_cases}),
   235                 [Rule_Cases.case_names [Binding.name_of Rep_name], Induct.cases_pred full_name]),
   159                   [Rule_Cases.case_names [Binding.name_of Rep_name], Induct.cases_pred full_name]),
   236               make @{thm type_definition.Rep_cases})
   160                 ((Binding.suffix_name "_cases" Abs_name, make @{thm type_definition.Abs_cases}),
   237           ||>> note_qualify ((Binding.suffix_name "_cases" Abs_name,
   161                   [Rule_Cases.case_names [Binding.name_of Abs_name], Induct.cases_type full_tname]),
   238                 [Rule_Cases.case_names [Binding.name_of Abs_name], Induct.cases_type full_tname]),
   162                 ((Binding.suffix_name "_induct" Rep_name, make @{thm type_definition.Rep_induct}),
   239               make @{thm type_definition.Abs_cases})
   163                   [Rule_Cases.case_names [Binding.name_of Rep_name], Induct.induct_pred full_name]),
   240           ||>> note_qualify ((Binding.suffix_name "_induct" Rep_name,
   164                 ((Binding.suffix_name "_induct" Abs_name, make @{thm type_definition.Abs_induct}),
   241                 [Rule_Cases.case_names [Binding.name_of Rep_name], Induct.induct_pred full_name]),
   165                   [Rule_Cases.case_names [Binding.name_of Abs_name], Induct.induct_type full_tname])]
   242               make @{thm type_definition.Rep_induct})
   166             ||> Sign.restore_naming thy1;
   243           ||>> note_qualify ((Binding.suffix_name "_induct" Abs_name,
   167           val info = {rep_type = oldT, abs_type = newT,
   244                 [Rule_Cases.case_names [Binding.name_of Abs_name], Induct.induct_type full_tname]),
   168             Rep_name = full Rep_name, Abs_name = full Abs_name,
   245               make @{thm type_definition.Abs_induct});
   169               inhabited = inhabited, type_definition = type_definition, set_def = set_def,
   246 
   170               Rep = Rep, Rep_inverse = Rep_inverse, Abs_inverse = Abs_inverse,
   247         val info = {rep_type = oldT, abs_type = newT,
   171               Rep_inject = Rep_inject, Abs_inject = Abs_inject, Rep_cases = Rep_cases,
   248           Rep_name = #1 (Term.dest_Const RepC), Abs_name = #1 (Term.dest_Const AbsC),
   172             Abs_cases = Abs_cases, Rep_induct = Rep_induct, Abs_induct = Abs_induct};
   249             inhabited = inhabited, type_definition = type_definition, set_def = set_def,
   173         in
   250             Rep = Rep, Rep_inverse = Rep_inverse, Abs_inverse = Abs_inverse,
   174           thy2
   251             Rep_inject = Rep_inject, Abs_inject = Abs_inject, Rep_cases = Rep_cases,
   175           |> put_info full_tname info
   252           Abs_cases = Abs_cases, Rep_induct = Rep_induct, Abs_induct = Abs_induct};
   176           |> Typedef_Interpretation.data full_tname
   253       in
   177           |> pair (full_tname, info)
   254         lthy2
   178         end);
   255         |> Local_Theory.declaration true (fn phi => put_info full_tname (transform_info phi info))
   179 
   256         |> Local_Theory.theory (Typedef_Interpretation.data full_tname)
   180 
   257         |> pair (full_tname, info)
   181     (* errors *)
   258       end;
   182 
   259 
   183     fun show_names pairs = commas_quote (map fst pairs);
   260   in ((goal, goal_pat, typedef_result), alias_lthy) end
   184 
       
   185     val illegal_vars =
       
   186       if null (Term.add_vars set []) andalso null (Term.add_tvars set []) then []
       
   187       else ["Illegal schematic variable(s) on rhs"];
       
   188 
       
   189     val dup_lhs_tfrees =
       
   190       (case duplicates (op =) lhs_tfrees of [] => []
       
   191       | dups => ["Duplicate type variables on lhs: " ^ show_names dups]);
       
   192 
       
   193     val extra_rhs_tfrees =
       
   194       (case fold (remove (op =)) lhs_tfrees rhs_tfrees of [] => []
       
   195       | extras => ["Extra type variables on rhs: " ^ show_names extras]);
       
   196 
       
   197     val illegal_frees =
       
   198       (case Term.add_frees set [] of [] => []
       
   199       | xs => ["Illegal variables on rhs: " ^ show_names xs]);
       
   200 
       
   201     val errs = illegal_vars @ dup_lhs_tfrees @ extra_rhs_tfrees @ illegal_frees;
       
   202     val _ = if null errs then () else error (cat_lines errs);
       
   203 
       
   204     (*test theory errors now!*)
       
   205     val test_thy = Theory.copy thy;
       
   206     val _ = typedef_result (Skip_Proof.make_thm test_thy goal) test_thy;
       
   207 
       
   208   in (set, goal, goal_pat, typedef_result) end
       
   209   handle ERROR msg =>
   261   handle ERROR msg =>
   210     cat_error msg ("The error(s) above occurred in typedef " ^ quote (Binding.str_of name));
   262     cat_error msg ("The error(s) above occurred in typedef " ^ quote (Binding.str_of name));
   211 
   263 
   212 
   264 
   213 (* add_typedef: tactic interface *)
   265 (* add_typedef: tactic interface *)
   214 
   266 
   215 fun add_typedef def opt_name typ set opt_morphs tac thy =
   267 fun add_typedef def opt_name typ set opt_morphs tac lthy =
   216   let
   268   let
   217     val name = the_default (#1 typ) opt_name;
   269     val name = the_default (#1 typ) opt_name;
   218     val (set, goal, _, typedef_result) =
   270     val ((goal, _, typedef_result), lthy') =
   219       prepare_typedef Syntax.check_term def name typ set opt_morphs thy;
   271       prepare_typedef Syntax.check_term def name typ set opt_morphs lthy;
   220     val inhabited = Goal.prove_global thy [] [] goal (K tac)
   272     val inhabited =
   221       handle ERROR msg => cat_error msg
   273       Goal.prove lthy' [] [] goal (K tac)
   222         ("Failed to prove non-emptiness of " ^ quote (Syntax.string_of_term_global thy set));
   274       |> Goal.norm_result |> Thm.close_derivation;
   223   in typedef_result inhabited thy end;
   275   in typedef_result inhabited lthy' end;
       
   276 
       
   277 fun add_typedef_global def opt_name typ set opt_morphs tac =
       
   278   Theory_Target.init NONE
       
   279   #> add_typedef def opt_name typ set opt_morphs tac
       
   280   #> Local_Theory.exit_result_global (apsnd o transform_info);
   224 
   281 
   225 
   282 
   226 (* typedef: proof interface *)
   283 (* typedef: proof interface *)
   227 
   284 
   228 local
   285 local
   229 
   286 
   230 fun gen_typedef prep_term ((def, name), typ, set, opt_morphs) thy =
   287 fun gen_typedef prep_term ((def, name), typ, set, opt_morphs) lthy =
   231   let
   288   let
   232     val (_, goal, goal_pat, typedef_result) =
   289     val ((goal, goal_pat, typedef_result), lthy') =
   233       prepare_typedef prep_term def name typ set opt_morphs thy;
   290       prepare_typedef prep_term def name typ set opt_morphs lthy;
   234     fun after_qed [[th]] = ProofContext.theory (snd o typedef_result th);
   291     fun after_qed [[th]] = snd o typedef_result th;
   235   in Proof.theorem_i NONE after_qed [[(goal, [goal_pat])]] (ProofContext.init thy) end;
   292   in Proof.theorem_i NONE after_qed [[(goal, [goal_pat])]] lthy' end;
   236 
   293 
   237 in
   294 in
   238 
   295 
   239 val typedef = gen_typedef Syntax.check_term;
   296 val typedef = gen_typedef Syntax.check_term;
   240 val typedef_cmd = gen_typedef Syntax.read_term;
   297 val typedef_cmd = gen_typedef Syntax.read_term;
   248 local structure P = OuterParse in
   305 local structure P = OuterParse in
   249 
   306 
   250 val _ = OuterKeyword.keyword "morphisms";
   307 val _ = OuterKeyword.keyword "morphisms";
   251 
   308 
   252 val _ =
   309 val _ =
   253   OuterSyntax.command "typedef" "HOL type definition (requires non-emptiness proof)"
   310   OuterSyntax.local_theory_to_proof "typedef" "HOL type definition (requires non-emptiness proof)"
   254     OuterKeyword.thy_goal
   311     OuterKeyword.thy_goal
   255     (Scan.optional (P.$$$ "(" |--
   312     (Scan.optional (P.$$$ "(" |--
   256         ((P.$$$ "open" >> K false) -- Scan.option P.binding ||
   313         ((P.$$$ "open" >> K false) -- Scan.option P.binding ||
   257           P.binding >> (fn s => (true, SOME s))) --| P.$$$ ")") (true, NONE) --
   314           P.binding >> (fn s => (true, SOME s))) --| P.$$$ ")") (true, NONE) --
   258       (P.type_args -- P.binding) -- P.opt_mixfix -- (P.$$$ "=" |-- P.term) --
   315       (P.type_args -- P.binding) -- P.opt_mixfix -- (P.$$$ "=" |-- P.term) --
   259       Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding))
   316       Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding))
   260     >> (fn ((((((def, opt_name), (vs, t)), mx), A), morphs)) =>
   317     >> (fn ((((((def, opt_name), (vs, t)), mx), A), morphs)) =>
   261         Toplevel.print o Toplevel.theory_to_proof
   318         typedef_cmd ((def, the_default t opt_name), (t, vs, mx), A, morphs)));
   262           (typedef_cmd ((def, the_default t opt_name), (t, vs, mx), A, morphs))));
       
   263 
   319 
   264 end;
   320 end;
   265 
   321 
   266 val setup = Typedef_Interpretation.init;
       
   267 
       
   268 end;
   322 end;
       
   323