| author | wenzelm | 
| Sat, 16 Jul 2016 00:38:33 +0200 | |
| changeset 63513 | 9f8d06f23c09 | 
| parent 61261 | ddb2da7cb2e4 | 
| child 69732 | 49d25343d3d4 | 
| permissions | -rw-r--r-- | 
| 35626 | 1 | (* Title: Pure/Isar/typedecl.ML | 
| 2 | Author: Makarius | |
| 3 | ||
| 36172 | 4 | Type declarations (with object-logic arities) and type abbreviations. | 
| 35626 | 5 | *) | 
| 6 | ||
| 7 | signature TYPEDECL = | |
| 8 | sig | |
| 35838 | 9 | val read_constraint: Proof.context -> string option -> sort | 
| 61259 | 10 |   val basic_typedecl: {final: bool} -> binding * int * mixfix ->
 | 
| 11 | local_theory -> string * local_theory | |
| 12 |   val typedecl: {final: bool} -> binding * (string * sort) list * mixfix ->
 | |
| 13 | local_theory -> typ * local_theory | |
| 14 |   val typedecl_global: {final: bool} -> binding * (string * sort) list * mixfix ->
 | |
| 15 | theory -> typ * theory | |
| 36172 | 16 | val abbrev: binding * string list * mixfix -> typ -> local_theory -> string * local_theory | 
| 17 | val abbrev_cmd: binding * string list * mixfix -> string -> local_theory -> string * local_theory | |
| 18 | val abbrev_global: binding * string list * mixfix -> typ -> theory -> string * theory | |
| 35626 | 19 | end; | 
| 20 | ||
| 21 | structure Typedecl: TYPEDECL = | |
| 22 | struct | |
| 23 | ||
| 36153 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
35838diff
changeset | 24 | (* constraints *) | 
| 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
35838diff
changeset | 25 | |
| 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
35838diff
changeset | 26 | fun read_constraint _ NONE = dummyS | 
| 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
35838diff
changeset | 27 | | read_constraint ctxt (SOME s) = Syntax.read_sort ctxt s; | 
| 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
35838diff
changeset | 28 | |
| 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
35838diff
changeset | 29 | |
| 35834 
0c71e0d72d7a
eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
 wenzelm parents: 
35740diff
changeset | 30 | (* primitives *) | 
| 
0c71e0d72d7a
eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
 wenzelm parents: 
35740diff
changeset | 31 | |
| 36172 | 32 | fun basic_decl decl (b, n, mx) lthy = | 
| 35834 
0c71e0d72d7a
eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
 wenzelm parents: 
35740diff
changeset | 33 | let val name = Local_Theory.full_name lthy b in | 
| 
0c71e0d72d7a
eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
 wenzelm parents: 
35740diff
changeset | 34 | lthy | 
| 38757 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 wenzelm parents: 
38388diff
changeset | 35 | |> Local_Theory.background_theory (decl name) | 
| 35834 
0c71e0d72d7a
eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
 wenzelm parents: 
35740diff
changeset | 36 | |> Local_Theory.type_notation true Syntax.mode_default [(Type (name, replicate n dummyT), mx)] | 
| 
0c71e0d72d7a
eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
 wenzelm parents: 
35740diff
changeset | 37 | |> Local_Theory.type_alias b name | 
| 
0c71e0d72d7a
eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
 wenzelm parents: 
35740diff
changeset | 38 | |> pair name | 
| 
0c71e0d72d7a
eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
 wenzelm parents: 
35740diff
changeset | 39 | end; | 
| 
0c71e0d72d7a
eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
 wenzelm parents: 
35740diff
changeset | 40 | |
| 
0c71e0d72d7a
eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
 wenzelm parents: 
35740diff
changeset | 41 | |
| 36172 | 42 | (* global type -- without dependencies on type parameters of the context *) | 
| 43 | ||
| 44 | fun global_type lthy (b, raw_args) = | |
| 35626 | 45 | let | 
| 42381 
309ec68442c6
added Binding.print convenience, which includes quote already;
 wenzelm parents: 
42375diff
changeset | 46 | fun err msg = error (msg ^ " in type declaration " ^ Binding.print b); | 
| 35681 | 47 | |
| 35834 
0c71e0d72d7a
eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
 wenzelm parents: 
35740diff
changeset | 48 | val _ = has_duplicates (eq_fst op =) raw_args andalso err "Duplicate parameters"; | 
| 42360 | 49 | val args = map (TFree o Proof_Context.check_tfree lthy) raw_args; | 
| 36156 | 50 | val T = Type (Local_Theory.full_name lthy b, args); | 
| 35681 | 51 | |
| 52 | val bad_args = | |
| 53 | #2 (Term.dest_Type (Logic.type_map (singleton (Variable.polymorphic lthy)) T)) | |
| 54 | |> filter_out Term.is_TVar; | |
| 36172 | 55 | val _ = null bad_args orelse | 
| 35740 
d3726291f252
added typedecl_wrt, which affects default sorts of type args;
 wenzelm parents: 
35714diff
changeset | 56 |       err ("Locally fixed type arguments " ^
 | 
| 35834 
0c71e0d72d7a
eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
 wenzelm parents: 
35740diff
changeset | 57 | commas_quote (map (Syntax.string_of_typ lthy) bad_args)); | 
| 36172 | 58 | in T end; | 
| 59 | ||
| 61250 
2f77019f6d0a
clarified deps entry: global names for arguments;
 wenzelm parents: 
61247diff
changeset | 60 | fun final_type (b, n) lthy = | 
| 61246 
077b88f9ec16
HOL typedef with explicit dependency checks according to Ondrey Kuncar, 07-Jul-2015, 16-Jul-2015, 30-Jul-2015;
 wenzelm parents: 
59970diff
changeset | 61 | let | 
| 61250 
2f77019f6d0a
clarified deps entry: global names for arguments;
 wenzelm parents: 
61247diff
changeset | 62 | val c = Local_Theory.full_name lthy b; | 
| 
2f77019f6d0a
clarified deps entry: global names for arguments;
 wenzelm parents: 
61247diff
changeset | 63 | val args = map (fn a => TFree (a, [])) (Name.invent Name.context Name.aT n); | 
| 61255 
15865e0c5598
eliminated separate type Theory.dep: use typeargs uniformly for consts/types;
 wenzelm parents: 
61250diff
changeset | 64 | in | 
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61259diff
changeset | 65 | Local_Theory.background_theory | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61259diff
changeset | 66 | (Theory.add_deps (Proof_Context.defs_context lthy) "" (Theory.type_dep (c, args)) []) lthy | 
| 61255 
15865e0c5598
eliminated separate type Theory.dep: use typeargs uniformly for consts/types;
 wenzelm parents: 
61250diff
changeset | 67 | end; | 
| 61250 
2f77019f6d0a
clarified deps entry: global names for arguments;
 wenzelm parents: 
61247diff
changeset | 68 | |
| 61259 | 69 | fun basic_typedecl {final} (b, n, mx) lthy =
 | 
| 61250 
2f77019f6d0a
clarified deps entry: global names for arguments;
 wenzelm parents: 
61247diff
changeset | 70 | lthy | 
| 
2f77019f6d0a
clarified deps entry: global names for arguments;
 wenzelm parents: 
61247diff
changeset | 71 | |> basic_decl (fn name => | 
| 
2f77019f6d0a
clarified deps entry: global names for arguments;
 wenzelm parents: 
61247diff
changeset | 72 | Sign.add_type lthy (b, n, NoSyn) #> | 
| 
2f77019f6d0a
clarified deps entry: global names for arguments;
 wenzelm parents: 
61247diff
changeset | 73 | (case Object_Logic.get_base_sort lthy of | 
| 
2f77019f6d0a
clarified deps entry: global names for arguments;
 wenzelm parents: 
61247diff
changeset | 74 | SOME S => Axclass.arity_axiomatization (name, replicate n S, S) | 
| 
2f77019f6d0a
clarified deps entry: global names for arguments;
 wenzelm parents: 
61247diff
changeset | 75 | | NONE => I)) (b, n, mx) | 
| 
2f77019f6d0a
clarified deps entry: global names for arguments;
 wenzelm parents: 
61247diff
changeset | 76 | ||> final ? final_type (b, n); | 
| 61246 
077b88f9ec16
HOL typedef with explicit dependency checks according to Ondrey Kuncar, 07-Jul-2015, 16-Jul-2015, 30-Jul-2015;
 wenzelm parents: 
59970diff
changeset | 77 | |
| 36172 | 78 | |
| 79 | (* type declarations *) | |
| 80 | ||
| 61259 | 81 | fun typedecl {final} (b, raw_args, mx) lthy =
 | 
| 36172 | 82 | let val T = global_type lthy (b, raw_args) in | 
| 35681 | 83 | lthy | 
| 61259 | 84 |     |> basic_typedecl {final = final} (b, length raw_args, mx)
 | 
| 35834 
0c71e0d72d7a
eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
 wenzelm parents: 
35740diff
changeset | 85 | |> snd | 
| 35681 | 86 | |> Variable.declare_typ T | 
| 35626 | 87 | |> pair T | 
| 88 | end; | |
| 89 | ||
| 61259 | 90 | fun typedecl_global {final} decl =
 | 
| 38388 | 91 | Named_Target.theory_init | 
| 61259 | 92 |   #> typedecl {final = final} decl
 | 
| 35681 | 93 | #> Local_Theory.exit_result_global Morphism.typ; | 
| 94 | ||
| 36172 | 95 | |
| 96 | (* type abbreviations *) | |
| 97 | ||
| 36457 
7355af2a7e8a
tuned user-level type abbrevs: explicit warning concerning ignored sort constraints -- sorts never affect formation of types and type abbrevs strip sorts internally;
 wenzelm parents: 
36179diff
changeset | 98 | local | 
| 
7355af2a7e8a
tuned user-level type abbrevs: explicit warning concerning ignored sort constraints -- sorts never affect formation of types and type abbrevs strip sorts internally;
 wenzelm parents: 
36179diff
changeset | 99 | |
| 36172 | 100 | fun gen_abbrev prep_typ (b, vs, mx) raw_rhs lthy = | 
| 101 | let | |
| 102 | val Type (name, _) = global_type lthy (b, map (rpair dummyS) vs); | |
| 36457 
7355af2a7e8a
tuned user-level type abbrevs: explicit warning concerning ignored sort constraints -- sorts never affect formation of types and type abbrevs strip sorts internally;
 wenzelm parents: 
36179diff
changeset | 103 | val rhs = prep_typ b lthy raw_rhs | 
| 42381 
309ec68442c6
added Binding.print convenience, which includes quote already;
 wenzelm parents: 
42375diff
changeset | 104 |       handle ERROR msg => cat_error msg ("in type abbreviation " ^ Binding.print b);
 | 
| 36172 | 105 | in | 
| 106 | lthy | |
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42360diff
changeset | 107 | |> basic_decl (fn _ => Sign.add_type_abbrev lthy (b, vs, rhs)) (b, length vs, mx) | 
| 36172 | 108 | |> snd | 
| 109 | |> pair name | |
| 110 | end; | |
| 111 | ||
| 36457 
7355af2a7e8a
tuned user-level type abbrevs: explicit warning concerning ignored sort constraints -- sorts never affect formation of types and type abbrevs strip sorts internally;
 wenzelm parents: 
36179diff
changeset | 112 | fun read_abbrev b ctxt raw_rhs = | 
| 
7355af2a7e8a
tuned user-level type abbrevs: explicit warning concerning ignored sort constraints -- sorts never affect formation of types and type abbrevs strip sorts internally;
 wenzelm parents: 
36179diff
changeset | 113 | let | 
| 42360 | 114 | val rhs = Proof_Context.read_typ_syntax (ctxt |> Proof_Context.set_defsort []) raw_rhs; | 
| 36457 
7355af2a7e8a
tuned user-level type abbrevs: explicit warning concerning ignored sort constraints -- sorts never affect formation of types and type abbrevs strip sorts internally;
 wenzelm parents: 
36179diff
changeset | 115 | val ignored = Term.fold_atyps_sorts (fn (_, []) => I | (T, _) => insert (op =) T) rhs []; | 
| 
7355af2a7e8a
tuned user-level type abbrevs: explicit warning concerning ignored sort constraints -- sorts never affect formation of types and type abbrevs strip sorts internally;
 wenzelm parents: 
36179diff
changeset | 116 | val _ = | 
| 56294 
85911b8a6868
prefer Context_Position where a context is available;
 wenzelm parents: 
51685diff
changeset | 117 | if not (null ignored) andalso Context_Position.is_visible ctxt then | 
| 
85911b8a6868
prefer Context_Position where a context is available;
 wenzelm parents: 
51685diff
changeset | 118 | warning | 
| 
85911b8a6868
prefer Context_Position where a context is available;
 wenzelm parents: 
51685diff
changeset | 119 |           ("Ignoring sort constraints in type variables(s): " ^
 | 
| 
85911b8a6868
prefer Context_Position where a context is available;
 wenzelm parents: 
51685diff
changeset | 120 | commas_quote (map (Syntax.string_of_typ ctxt) (rev ignored)) ^ | 
| 
85911b8a6868
prefer Context_Position where a context is available;
 wenzelm parents: 
51685diff
changeset | 121 | "\nin type abbreviation " ^ Binding.print b) | 
| 
85911b8a6868
prefer Context_Position where a context is available;
 wenzelm parents: 
51685diff
changeset | 122 | else (); | 
| 36457 
7355af2a7e8a
tuned user-level type abbrevs: explicit warning concerning ignored sort constraints -- sorts never affect formation of types and type abbrevs strip sorts internally;
 wenzelm parents: 
36179diff
changeset | 123 | in rhs end; | 
| 
7355af2a7e8a
tuned user-level type abbrevs: explicit warning concerning ignored sort constraints -- sorts never affect formation of types and type abbrevs strip sorts internally;
 wenzelm parents: 
36179diff
changeset | 124 | |
| 
7355af2a7e8a
tuned user-level type abbrevs: explicit warning concerning ignored sort constraints -- sorts never affect formation of types and type abbrevs strip sorts internally;
 wenzelm parents: 
36179diff
changeset | 125 | in | 
| 
7355af2a7e8a
tuned user-level type abbrevs: explicit warning concerning ignored sort constraints -- sorts never affect formation of types and type abbrevs strip sorts internally;
 wenzelm parents: 
36179diff
changeset | 126 | |
| 42360 | 127 | val abbrev = gen_abbrev (K Proof_Context.cert_typ_syntax); | 
| 36457 
7355af2a7e8a
tuned user-level type abbrevs: explicit warning concerning ignored sort constraints -- sorts never affect formation of types and type abbrevs strip sorts internally;
 wenzelm parents: 
36179diff
changeset | 128 | val abbrev_cmd = gen_abbrev read_abbrev; | 
| 
7355af2a7e8a
tuned user-level type abbrevs: explicit warning concerning ignored sort constraints -- sorts never affect formation of types and type abbrevs strip sorts internally;
 wenzelm parents: 
36179diff
changeset | 129 | |
| 
7355af2a7e8a
tuned user-level type abbrevs: explicit warning concerning ignored sort constraints -- sorts never affect formation of types and type abbrevs strip sorts internally;
 wenzelm parents: 
36179diff
changeset | 130 | end; | 
| 36172 | 131 | |
| 132 | fun abbrev_global decl rhs = | |
| 38388 | 133 | Named_Target.theory_init | 
| 36172 | 134 | #> abbrev decl rhs | 
| 135 | #> Local_Theory.exit_result_global (K I); | |
| 136 | ||
| 35626 | 137 | end; |