src/Pure/Isar/typedecl.ML
author haftmann
Wed, 23 Jan 2019 17:54:50 +0000
changeset 69732 49d25343d3d4
parent 61261 ddb2da7cb2e4
permissions -rw-r--r--
combinator to lift local theory update to theory update
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35626
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/typedecl.ML
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
     3
36172
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
     4
Type declarations (with object-logic arities) and type abbreviations.
35626
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
     5
*)
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
     6
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
     7
signature TYPEDECL =
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
     8
sig
35838
c8bd075c4de8 support type arguments with sort constraints;
wenzelm
parents: 35834
diff changeset
     9
  val read_constraint: Proof.context -> string option -> sort
61259
6dc3d5d50e57 tuned signature;
wenzelm
parents: 61255
diff changeset
    10
  val basic_typedecl: {final: bool} -> binding * int * mixfix ->
6dc3d5d50e57 tuned signature;
wenzelm
parents: 61255
diff changeset
    11
    local_theory -> string * local_theory
6dc3d5d50e57 tuned signature;
wenzelm
parents: 61255
diff changeset
    12
  val typedecl: {final: bool} -> binding * (string * sort) list * mixfix ->
6dc3d5d50e57 tuned signature;
wenzelm
parents: 61255
diff changeset
    13
    local_theory -> typ * local_theory
6dc3d5d50e57 tuned signature;
wenzelm
parents: 61255
diff changeset
    14
  val typedecl_global: {final: bool} -> binding * (string * sort) list * mixfix ->
6dc3d5d50e57 tuned signature;
wenzelm
parents: 61255
diff changeset
    15
    theory -> typ * theory
36172
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    16
  val abbrev: binding * string list * mixfix -> typ -> local_theory -> string * local_theory
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    17
  val abbrev_cmd: binding * string list * mixfix -> string -> local_theory -> string * local_theory
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    18
  val abbrev_global: binding * string list * mixfix -> typ -> theory -> string * theory
35626
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
    19
end;
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
    20
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
    21
structure Typedecl: TYPEDECL =
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
    22
struct
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
    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: 35838
diff 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: 35838
diff 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: 35838
diff 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: 35838
diff 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: 35838
diff 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: 35838
diff changeset
    29
35834
0c71e0d72d7a eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
wenzelm
parents: 35740
diff changeset
    30
(* primitives *)
0c71e0d72d7a eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
wenzelm
parents: 35740
diff changeset
    31
36172
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    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: 35740
diff 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: 35740
diff 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: 38388
diff 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: 35740
diff 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: 35740
diff 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: 35740
diff changeset
    38
    |> pair name
0c71e0d72d7a eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
wenzelm
parents: 35740
diff changeset
    39
  end;
0c71e0d72d7a eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
wenzelm
parents: 35740
diff changeset
    40
0c71e0d72d7a eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
wenzelm
parents: 35740
diff changeset
    41
36172
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    42
(* global type -- without dependencies on type parameters of the context *)
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    43
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    44
fun global_type lthy (b, raw_args) =
35626
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
    45
  let
42381
309ec68442c6 added Binding.print convenience, which includes quote already;
wenzelm
parents: 42375
diff changeset
    46
    fun err msg = error (msg ^ " in type declaration " ^ Binding.print b);
35681
8b22a498b034 localized typedecl;
wenzelm
parents: 35626
diff changeset
    47
35834
0c71e0d72d7a eliminated slightly odd typedecl_wrt in favour of explicit predeclare_constraints (which also works for recursive types);
wenzelm
parents: 35740
diff changeset
    48
    val _ = has_duplicates (eq_fst op =) raw_args andalso err "Duplicate parameters";
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 38831
diff changeset
    49
    val args = map (TFree o Proof_Context.check_tfree lthy) raw_args;
36156
6f0a8f6b1671 explicit ProofContext.check_tfree;
wenzelm
parents: 36153
diff changeset
    50
    val T = Type (Local_Theory.full_name lthy b, args);
35681
8b22a498b034 localized typedecl;
wenzelm
parents: 35626
diff changeset
    51
8b22a498b034 localized typedecl;
wenzelm
parents: 35626
diff changeset
    52
    val bad_args =
8b22a498b034 localized typedecl;
wenzelm
parents: 35626
diff changeset
    53
      #2 (Term.dest_Type (Logic.type_map (singleton (Variable.polymorphic lthy)) T))
8b22a498b034 localized typedecl;
wenzelm
parents: 35626
diff changeset
    54
      |> filter_out Term.is_TVar;
36172
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    55
    val _ = null bad_args orelse
35740
d3726291f252 added typedecl_wrt, which affects default sorts of type args;
wenzelm
parents: 35714
diff 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: 35740
diff changeset
    57
        commas_quote (map (Syntax.string_of_typ lthy) bad_args));
36172
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    58
  in T end;
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    59
61250
2f77019f6d0a clarified deps entry: global names for arguments;
wenzelm
parents: 61247
diff 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: 59970
diff changeset
    61
  let
61250
2f77019f6d0a clarified deps entry: global names for arguments;
wenzelm
parents: 61247
diff changeset
    62
    val c = Local_Theory.full_name lthy b;
2f77019f6d0a clarified deps entry: global names for arguments;
wenzelm
parents: 61247
diff 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: 61250
diff changeset
    64
  in
61261
ddb2da7cb2e4 more explicit Defs.context: use proper name spaces as far as possible;
wenzelm
parents: 61259
diff changeset
    65
    Local_Theory.background_theory
ddb2da7cb2e4 more explicit Defs.context: use proper name spaces as far as possible;
wenzelm
parents: 61259
diff 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: 61250
diff changeset
    67
  end;
61250
2f77019f6d0a clarified deps entry: global names for arguments;
wenzelm
parents: 61247
diff changeset
    68
61259
6dc3d5d50e57 tuned signature;
wenzelm
parents: 61255
diff changeset
    69
fun basic_typedecl {final} (b, n, mx) lthy =
61250
2f77019f6d0a clarified deps entry: global names for arguments;
wenzelm
parents: 61247
diff changeset
    70
  lthy
2f77019f6d0a clarified deps entry: global names for arguments;
wenzelm
parents: 61247
diff changeset
    71
  |> basic_decl (fn name =>
2f77019f6d0a clarified deps entry: global names for arguments;
wenzelm
parents: 61247
diff changeset
    72
    Sign.add_type lthy (b, n, NoSyn) #>
2f77019f6d0a clarified deps entry: global names for arguments;
wenzelm
parents: 61247
diff changeset
    73
    (case Object_Logic.get_base_sort lthy of
2f77019f6d0a clarified deps entry: global names for arguments;
wenzelm
parents: 61247
diff changeset
    74
      SOME S => Axclass.arity_axiomatization (name, replicate n S, S)
2f77019f6d0a clarified deps entry: global names for arguments;
wenzelm
parents: 61247
diff changeset
    75
    | NONE => I)) (b, n, mx)
2f77019f6d0a clarified deps entry: global names for arguments;
wenzelm
parents: 61247
diff 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: 59970
diff changeset
    77
36172
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    78
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    79
(* type declarations *)
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    80
61259
6dc3d5d50e57 tuned signature;
wenzelm
parents: 61255
diff changeset
    81
fun typedecl {final} (b, raw_args, mx) lthy =
36172
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    82
  let val T = global_type lthy (b, raw_args) in
35681
8b22a498b034 localized typedecl;
wenzelm
parents: 35626
diff changeset
    83
    lthy
61259
6dc3d5d50e57 tuned signature;
wenzelm
parents: 61255
diff changeset
    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: 35740
diff changeset
    85
    |> snd
35681
8b22a498b034 localized typedecl;
wenzelm
parents: 35626
diff changeset
    86
    |> Variable.declare_typ T
35626
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
    87
    |> pair T
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
    88
  end;
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
    89
61259
6dc3d5d50e57 tuned signature;
wenzelm
parents: 61255
diff changeset
    90
fun typedecl_global {final} decl =
69732
49d25343d3d4 combinator to lift local theory update to theory update
haftmann
parents: 61261
diff changeset
    91
  Named_Target.theory_map_result Morphism.typ (typedecl {final = final} decl);
35681
8b22a498b034 localized typedecl;
wenzelm
parents: 35626
diff changeset
    92
36172
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    93
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    94
(* type abbreviations *)
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    95
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: 36179
diff changeset
    96
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: 36179
diff changeset
    97
36172
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    98
fun gen_abbrev prep_typ (b, vs, mx) raw_rhs lthy =
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
    99
  let
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
   100
    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: 36179
diff changeset
   101
    val rhs = prep_typ b lthy raw_rhs
42381
309ec68442c6 added Binding.print convenience, which includes quote already;
wenzelm
parents: 42375
diff changeset
   102
      handle ERROR msg => cat_error msg ("in type abbreviation " ^ Binding.print b);
36172
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
   103
  in
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
   104
    lthy
42375
774df7c59508 report Name_Space.declare/define, relatively to context;
wenzelm
parents: 42360
diff changeset
   105
    |> basic_decl (fn _ => Sign.add_type_abbrev lthy (b, vs, rhs)) (b, length vs, mx)
36172
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
   106
    |> snd
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
   107
    |> pair name
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
   108
  end;
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
   109
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: 36179
diff changeset
   110
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: 36179
diff changeset
   111
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 38831
diff changeset
   112
    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: 36179
diff changeset
   113
    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: 36179
diff changeset
   114
    val _ =
56294
85911b8a6868 prefer Context_Position where a context is available;
wenzelm
parents: 51685
diff changeset
   115
      if not (null ignored) andalso Context_Position.is_visible ctxt then
85911b8a6868 prefer Context_Position where a context is available;
wenzelm
parents: 51685
diff changeset
   116
        warning
85911b8a6868 prefer Context_Position where a context is available;
wenzelm
parents: 51685
diff changeset
   117
          ("Ignoring sort constraints in type variables(s): " ^
85911b8a6868 prefer Context_Position where a context is available;
wenzelm
parents: 51685
diff changeset
   118
            commas_quote (map (Syntax.string_of_typ ctxt) (rev ignored)) ^
85911b8a6868 prefer Context_Position where a context is available;
wenzelm
parents: 51685
diff changeset
   119
            "\nin type abbreviation " ^ Binding.print b)
85911b8a6868 prefer Context_Position where a context is available;
wenzelm
parents: 51685
diff changeset
   120
      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: 36179
diff changeset
   121
  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: 36179
diff changeset
   122
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: 36179
diff changeset
   123
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: 36179
diff changeset
   124
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 38831
diff changeset
   125
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: 36179
diff changeset
   126
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: 36179
diff changeset
   127
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: 36179
diff changeset
   128
end;
36172
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
   129
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
   130
fun abbrev_global decl rhs =
69732
49d25343d3d4 combinator to lift local theory update to theory update
haftmann
parents: 61261
diff changeset
   131
  Named_Target.theory_map_result (K I) (abbrev decl rhs);
36172
fc407d02af4a local type abbreviations;
wenzelm
parents: 36156
diff changeset
   132
35626
06197484c6ad separate structure Typedecl;
wenzelm
parents:
diff changeset
   133
end;