src/Pure/Isar/code.ML
author haftmann
Sat, 19 Jul 2025 18:41:55 +0200
changeset 82886 8d1e295aab70
parent 82858 52cf8f3f1652
permissions -rw-r--r--
clarified name and status of auxiliary operation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     1
(*  Title:      Pure/Isar/code.ML
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     3
34173
458ced35abb8 reduced code generator cache to the baremost minimum
haftmann
parents: 33977
diff changeset
     4
Abstract executable ingredients of theory.  Management of data
458ced35abb8 reduced code generator cache to the baremost minimum
haftmann
parents: 33977
diff changeset
     5
dependent on executable ingredients as synchronized cache; purged
458ced35abb8 reduced code generator cache to the baremost minimum
haftmann
parents: 33977
diff changeset
     6
on any change of underlying executable ingredients.
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     7
*)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     8
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     9
signature CODE =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    10
sig
31957
a9742afd403e tuned interface of structure Code
haftmann
parents: 31890
diff changeset
    11
  (*constants*)
a9742afd403e tuned interface of structure Code
haftmann
parents: 31890
diff changeset
    12
  val check_const: theory -> term -> string
82858
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
    13
  val read_const: Proof.context -> string -> string
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
    14
  val string_of_const: theory -> string -> string
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
    15
  val args_number: theory -> string -> int
31957
a9742afd403e tuned interface of structure Code
haftmann
parents: 31890
diff changeset
    16
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
    17
  (*constructor sets*)
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
    18
  val constrset_of_consts: theory -> (string * typ) list
40726
16dcfedc4eb7 keep type variable arguments of datatype constructors in bookkeeping
haftmann
parents: 40564
diff changeset
    19
    -> string * ((string * sort) list * (string * ((string * sort) list * typ list)) list)
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
    20
34874
89c230bf8feb added code certificates
haftmann
parents: 34272
diff changeset
    21
  (*code equations and certificates*)
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
    22
  val assert_eqn: theory -> thm * bool -> thm * bool
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
    23
  val assert_abs_eqn: theory -> string option -> thm -> thm * (string * string)
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
    24
  type cert
34874
89c230bf8feb added code certificates
haftmann
parents: 34272
diff changeset
    25
  val constrain_cert: theory -> sort list -> cert -> cert
49971
8b50286c36d3 close code theorems explicitly after preprocessing
haftmann
parents: 49904
diff changeset
    26
  val conclude_cert: cert -> cert
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
    27
  val typargs_deps_of_cert: theory -> cert -> (string * sort) list * (string * typ list) list
36209
566be5d48090 more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
haftmann
parents: 36202
diff changeset
    28
  val equations_of_cert: theory -> cert -> ((string * sort) list * typ)
77702
haftmann
parents: 75399
diff changeset
    29
    * (((string option * term) list * (string option * term)) * (thm option * bool)) list option
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
    30
  val pretty_cert: theory -> cert -> Pretty.T list
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
    31
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
    32
  (*executable code*)
66328
cf9ce8016da1 clarified
haftmann
parents: 66324
diff changeset
    33
  type constructors
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
    34
  type abs_type
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
    35
  val type_interpretation: (string -> theory -> theory) -> theory -> theory
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
    36
  val datatype_interpretation: (string * constructors -> theory -> theory) -> theory -> theory
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
    37
  val abstype_interpretation: (string * abs_type -> theory -> theory) -> theory -> theory
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
    38
  val declare_datatype_global: (string * typ) list -> theory -> theory
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
    39
  val declare_datatype_cmd: string list -> theory -> theory
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
    40
  val declare_abstype: thm -> local_theory -> local_theory
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
    41
  val declare_abstype_global: thm -> theory -> theory
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
    42
  val declare_default_eqns: (thm * bool) list -> local_theory -> local_theory
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
    43
  val declare_default_eqns_global: (thm * bool) list -> theory -> theory
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
    44
  val declare_eqns: (thm * bool) list -> local_theory -> local_theory
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
    45
  val declare_eqns_global: (thm * bool) list -> theory -> theory
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
    46
  val add_eqn_global: thm * bool -> theory -> theory
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
    47
  val del_eqn_global: thm -> theory -> theory
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
    48
  val declare_abstract_eqn: thm -> local_theory -> local_theory
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
    49
  val declare_abstract_eqn_global: thm -> theory -> theory
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
    50
  val declare_aborting_global: string -> theory -> theory
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
    51
  val declare_unimplemented_global: string -> theory -> theory
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
    52
  val declare_case_global: thm -> theory -> theory
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
    53
  val declare_undefined_global: string -> theory -> theory
66328
cf9ce8016da1 clarified
haftmann
parents: 66324
diff changeset
    54
  val get_type: theory -> string -> constructors * bool
35299
4f4d5bf4ea08 proper distinction of code datatypes and abstypes
haftmann
parents: 35226
diff changeset
    55
  val get_type_of_constr_or_abstr: theory -> string -> (string * bool) option
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
    56
  val is_constr: theory -> string -> bool
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
    57
  val is_abstr: theory -> string -> bool
57429
4aef934d43ad tuned interface
haftmann
parents: 56811
diff changeset
    58
  val get_cert: Proof.context -> ((thm * bool) list -> (thm * bool) list option) list
4aef934d43ad tuned interface
haftmann
parents: 56811
diff changeset
    59
    -> string -> cert
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
    60
  type case_schema
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
    61
  val get_case_schema: theory -> string -> case_schema option
37438
4906ab970316 maintain cong rules for case combinators
haftmann
parents: 37425
diff changeset
    62
  val get_case_cong: theory -> string -> thm option
66189
23917e861eaa treat "undefined" constants internally as special form of case combinators
haftmann
parents: 66167
diff changeset
    63
  val is_undefined: theory -> string -> bool
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    64
  val print_codesetup: theory -> unit
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
    65
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
    66
  (*transitional*)
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
    67
  val only_single_equation: bool Config.T
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
    68
  val prepend_allowed: bool Config.T
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    69
end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    70
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    71
signature CODE_DATA_ARGS =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    72
sig
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    73
  type T
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    74
  val empty: T
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    75
end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    76
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    77
signature CODE_DATA =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    78
sig
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    79
  type T
39397
9b0a8d72edc8 ignore code cache optionally
haftmann
parents: 39134
diff changeset
    80
  val change: theory option -> (T -> T) -> T
9b0a8d72edc8 ignore code cache optionally
haftmann
parents: 39134
diff changeset
    81
  val change_yield: theory option -> (T -> 'a * T) -> 'a * T
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    82
end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    83
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    84
signature PRIVATE_CODE =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    85
sig
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    86
  include CODE
51368
2ea5c7c2d825 tuned signature -- prefer terminology of Scala and Axiom;
wenzelm
parents: 49971
diff changeset
    87
  val declare_data: Any.T -> serial
2ea5c7c2d825 tuned signature -- prefer terminology of Scala and Axiom;
wenzelm
parents: 49971
diff changeset
    88
  val change_yield_data: serial * ('a -> Any.T) * (Any.T -> 'a)
34251
cd642bb91f64 code cache only persists on equal theories
haftmann
parents: 34244
diff changeset
    89
    -> theory -> ('a -> 'b * 'a) -> 'b * 'a
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    90
end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    91
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    92
structure Code : PRIVATE_CODE =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    93
struct
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    94
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
    95
(** auxiliary **)
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
    96
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
    97
(* printing *)
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
    98
39134
917b4b6ba3d2 turned show_sorts/show_types into proper configuration options;
wenzelm
parents: 39020
diff changeset
    99
fun string_of_typ thy =
917b4b6ba3d2 turned show_sorts/show_types into proper configuration options;
wenzelm
parents: 39020
diff changeset
   100
  Syntax.string_of_typ (Config.put show_sorts true (Syntax.init_pretty_global thy));
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   101
42359
6ca5407863ed prefer local name spaces;
wenzelm
parents: 41493
diff changeset
   102
fun string_of_const thy c =
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   103
  let val ctxt = Proof_Context.init_global thy in
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51584
diff changeset
   104
    case Axclass.inst_of_param thy c of
42359
6ca5407863ed prefer local name spaces;
wenzelm
parents: 41493
diff changeset
   105
      SOME (c, tyco) =>
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   106
        Proof_Context.extern_const ctxt c ^ " " ^ enclose "[" "]"
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   107
          (Proof_Context.extern_type ctxt tyco)
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   108
    | NONE => Proof_Context.extern_const ctxt c
42359
6ca5407863ed prefer local name spaces;
wenzelm
parents: 41493
diff changeset
   109
  end;
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   110
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   111
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   112
(* transitional *)
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   113
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   114
val only_single_equation = Attrib.setup_config_bool \<^binding>\<open>code_only_single_equation\<close> (K false);
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   115
val prepend_allowed = Attrib.setup_config_bool \<^binding>\<open>code_prepend_allowed\<close> (K false);
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   116
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   117
val _ = Theory.setup (Theory.at_end ((fn thy => if Config.get_global thy prepend_allowed
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   118
  then thy |> Config.put_global prepend_allowed false |> SOME
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   119
  else NONE)));
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   120
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   121
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   122
(* constants *)
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   123
79411
700d4f16b5f2 minor performance tuning: proper Same.operation;
wenzelm
parents: 79232
diff changeset
   124
fun const_typ thy = Term.strip_sortsT o Sign.the_const_type thy;
49534
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   125
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   126
fun args_number thy = length o binder_types o const_typ thy;
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   127
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   128
fun devarify ty =
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   129
  let
75353
05f7f5454cb6 prefer build combinator
haftmann
parents: 74978
diff changeset
   130
    val tys = build (fold_atyps (fn TVar vi_sort => AList.update (op =) vi_sort) ty);
81507
08574da77b4a clarified signature: shorten common cases;
wenzelm
parents: 79411
diff changeset
   131
    val vs = Name.invent_global_types (length tys);
49534
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   132
    val mapping = map2 (fn v => fn (vi, sort) => (vi, TFree (v, sort))) vs tys;
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   133
  in Term.typ_subst_TVars mapping ty end;
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   134
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   135
fun typscheme thy (c, ty) =
79411
700d4f16b5f2 minor performance tuning: proper Same.operation;
wenzelm
parents: 79232
diff changeset
   136
  (map dest_TFree (Sign.const_typargs thy (c, ty)), Term.strip_sortsT ty);
49534
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   137
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   138
fun typscheme_equiv (ty1, ty2) =
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   139
  Type.raw_instance (devarify ty1, ty2) andalso Type.raw_instance (devarify ty2, ty1);
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   140
82858
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   141
local
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   142
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   143
fun check_bare_const thy t = case try dest_Const t
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   144
 of SOME c_ty => c_ty
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   145
  | NONE => error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   146
40362
82a066bff182 Code.check_const etc.: reject too specific types
haftmann
parents: 40316
diff changeset
   147
fun check_unoverload thy (c, ty) =
82a066bff182 Code.check_const etc.: reject too specific types
haftmann
parents: 40316
diff changeset
   148
  let
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51584
diff changeset
   149
    val c' = Axclass.unoverload_const thy (c, ty);
54604
1512fa5fe531 prefer sort-stripping const_typ over Sign.the_const_type whenever appropriate
haftmann
parents: 54603
diff changeset
   150
    val ty_decl = const_typ thy c';
45344
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 45211
diff changeset
   151
  in
49534
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   152
    if typscheme_equiv (ty_decl, Logic.varifyT_global ty)
45344
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 45211
diff changeset
   153
    then c'
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 45211
diff changeset
   154
    else
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 45211
diff changeset
   155
      error ("Type\n" ^ string_of_typ thy ty ^
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 45211
diff changeset
   156
        "\nof constant " ^ quote c ^
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 45211
diff changeset
   157
        "\nis too specific compared to declared type\n" ^
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 45211
diff changeset
   158
        string_of_typ thy ty_decl)
66167
1bd268ab885c more informative task_statistics;
wenzelm
parents: 66149
diff changeset
   159
  end;
40362
82a066bff182 Code.check_const etc.: reject too specific types
haftmann
parents: 40316
diff changeset
   160
82858
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   161
fun export_global ctxt thy =
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   162
  Variable.export_terms ctxt (Proof_Context.init_global thy);
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   163
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   164
in
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   165
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   166
fun check_const thy =
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   167
  check_bare_const thy #> check_unoverload thy;
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   168
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   169
fun read_bare_const thy =
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   170
  Syntax.read_term (Proof_Context.init_global thy) #> check_bare_const thy;
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   171
82858
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   172
fun read_const ctxt =
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   173
  let
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   174
    val thy = Proof_Context.theory_of ctxt
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   175
  in
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   176
    Syntax.read_term ctxt
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   177
    #> singleton (export_global ctxt thy)
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   178
    #> Logic.unvarify_types_global
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   179
    #> check_bare_const thy
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   180
    #> check_unoverload thy
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   181
  end;
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   182
82858
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   183
fun prospective_check_consts ctxt ts thy =
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   184
  ts
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   185
  |> export_global ctxt thy
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   186
  |> map Logic.unvarify_types_global
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   187
  |> map (check_const thy);
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   188
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   189
val prospective_const_args = Args.context :|--
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   190
  (fn ctxt => Scan.repeat1 Args.term >> prospective_check_consts ctxt);
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   191
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
   192
end;
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   193
32872
019201eb7e07 variables in type schemes must be renamed simultaneously with variables in equations
haftmann
parents: 32738
diff changeset
   194
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   195
(** executable specifications **)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   196
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   197
(* types *)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   198
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   199
datatype type_spec = Constructors of {
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   200
      constructors: (string * ((string * sort) list * typ list)) list,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   201
      case_combinators: string list}
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   202
  | Abstractor of {
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   203
      abs_rep: thm,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   204
      abstractor: string * ((string * sort) list * typ),
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   205
      projection: string,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   206
      more_abstract_functions: string list};
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   207
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   208
fun concrete_constructors_of (Constructors {constructors, ...}) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   209
      constructors
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   210
  | concrete_constructors_of _ =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   211
      [];
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   212
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   213
fun constructors_of (Constructors {constructors, ...}) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   214
      (constructors, false)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   215
  | constructors_of (Abstractor {abstractor = (co, (vs, ty)), ...}) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   216
      ([(co, (vs, [ty]))], true);
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   217
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   218
fun case_combinators_of (Constructors {case_combinators, ...}) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   219
      case_combinators
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   220
  | case_combinators_of (Abstractor _) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   221
      [];
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   222
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   223
fun add_case_combinator c (vs, Constructors {constructors, case_combinators}) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   224
  (vs, Constructors {constructors = constructors,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   225
    case_combinators = insert (op =) c case_combinators});
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   226
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   227
fun projection_of (Constructors _) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   228
      NONE
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   229
  | projection_of (Abstractor {projection, ...}) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   230
      SOME projection;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   231
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   232
fun abstract_functions_of (Constructors _) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   233
      []
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   234
  | abstract_functions_of (Abstractor {more_abstract_functions, projection, ...}) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   235
      projection :: more_abstract_functions;
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   236
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   237
fun add_abstract_function c (vs, Abstractor {abs_rep, abstractor, projection, more_abstract_functions}) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   238
  (vs, Abstractor {abs_rep = abs_rep, abstractor = abstractor, projection = projection,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   239
    more_abstract_functions = insert (op =) c more_abstract_functions});
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   240
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   241
fun join_same_types' (Constructors {constructors, case_combinators = case_combinators1},
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   242
    Constructors {case_combinators = case_combinators2, ...}) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   243
      Constructors {constructors = constructors,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   244
        case_combinators = merge (op =) (case_combinators1, case_combinators2)}
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   245
  | join_same_types' (Abstractor {abs_rep, abstractor, projection, more_abstract_functions = more_abstract_functions1},
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   246
      Abstractor {more_abstract_functions = more_abstract_functions2, ...}) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   247
      Abstractor {abs_rep = abs_rep, abstractor = abstractor, projection = projection,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   248
        more_abstract_functions = merge (op =) (more_abstract_functions1, more_abstract_functions2)};
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   249
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   250
fun join_same_types ((vs, spec1), (_, spec2)) = (vs, join_same_types' (spec1, spec2));
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   251
66125
28e782dd0278 tuned whitespace
haftmann
parents: 66024
diff changeset
   252
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   253
(* functions *)
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   254
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   255
datatype fun_spec =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   256
    Eqns of bool * (thm * bool) list
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   257
  | Proj of term * (string * string)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   258
  | Abstr of thm * (string * string);
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   259
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   260
val unimplemented = Eqns (true, []);
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   261
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   262
fun is_unimplemented (Eqns (true, [])) = true
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   263
  | is_unimplemented _ = false;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   264
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   265
fun is_default (Eqns (true, _)) = true
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   266
  | is_default _ = false;
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   267
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   268
val aborting = Eqns (false, []);
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   269
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   270
fun associated_abstype (Proj (_, tyco_abs)) = SOME tyco_abs
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   271
  | associated_abstype (Abstr (_, tyco_abs)) = SOME tyco_abs
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   272
  | associated_abstype _ = NONE;
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   273
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   274
66189
23917e861eaa treat "undefined" constants internally as special form of case combinators
haftmann
parents: 66167
diff changeset
   275
(* cases *)
23917e861eaa treat "undefined" constants internally as special form of case combinators
haftmann
parents: 66167
diff changeset
   276
75399
cdf84288d93c pass constructor arity as part of case certficiate
haftmann
parents: 75353
diff changeset
   277
type case_schema = int * (int * (string * int) option list);
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   278
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   279
datatype case_spec =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   280
    No_Case
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   281
  | Case of {schema: case_schema, tycos: string list, cong: thm}
66189
23917e861eaa treat "undefined" constants internally as special form of case combinators
haftmann
parents: 66167
diff changeset
   282
  | Undefined;
23917e861eaa treat "undefined" constants internally as special form of case combinators
haftmann
parents: 66167
diff changeset
   283
75399
cdf84288d93c pass constructor arity as part of case certficiate
haftmann
parents: 75353
diff changeset
   284
fun associated_datatypes (Case {tycos, schema = (_, (_, raw_cos)), ...}) = (tycos, map fst (map_filter I raw_cos))
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   285
  | associated_datatypes _ = ([], []);
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   286
66189
23917e861eaa treat "undefined" constants internally as special form of case combinators
haftmann
parents: 66167
diff changeset
   287
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   288
(** background theory data store **)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   289
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   290
(* historized declaration data *)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   291
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   292
structure History =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   293
struct
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   294
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   295
type 'a T = {
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   296
  entry: 'a,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   297
  suppressed: bool,     (*incompatible entries are merely suppressed after theory merge but sustain*)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   298
  history: serial list  (*explicit trace of declaration history supports non-monotonic declarations*)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   299
} Symtab.table;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   300
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   301
fun some_entry (SOME {suppressed = false, entry, ...}) = SOME entry
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   302
  | some_entry _ = NONE;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   303
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   304
fun lookup table =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   305
  Symtab.lookup table #> some_entry;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   306
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   307
fun register key entry table =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   308
  if is_some (Symtab.lookup table key)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   309
  then Symtab.map_entry key
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   310
    (fn {history, ...} => {entry = entry, suppressed = false, history = serial () :: history}) table
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   311
  else Symtab.update (key, {entry = entry, suppressed = false, history = [serial ()]}) table;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   312
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   313
fun modify_entry key f = Symtab.map_entry key
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   314
  (fn {entry, suppressed, history} => {entry = f entry, suppressed = suppressed, history = history});
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   315
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   316
fun all table = Symtab.dest table
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   317
  |> map_filter (fn (key, {entry, suppressed = false, ...}) => SOME (key, entry) | _ => NONE);
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   318
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   319
local
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   320
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   321
fun merge_history join_same
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   322
    ({entry = entry1, history = history1, ...}, {entry = entry2, history = history2, ...}) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   323
  let
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   324
    val history = merge (op =) (history1, history2);
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   325
    val entry = if hd history1 = hd history2 then join_same (entry1, entry2)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   326
      else if hd history = hd history1 then entry1 else entry2;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   327
  in {entry = entry, suppressed = false, history = history} end;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   328
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   329
in
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   330
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   331
fun join join_same tables = Symtab.join (K (merge_history join_same)) tables;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   332
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   333
fun suppress key = Symtab.map_entry key
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   334
  (fn {entry, history, ...} => {entry = entry, suppressed = true, history = history});
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   335
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   336
fun suppress_except f = Symtab.map (fn key => fn {entry, suppressed, history} =>
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   337
  {entry = entry, suppressed = suppressed orelse (not o f) (key, entry), history = history});
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   338
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   339
end;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   340
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   341
end;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   342
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   343
datatype specs = Specs of {
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   344
  types: ((string * sort) list * type_spec) History.T,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   345
  pending_eqns: (thm * bool) list Symtab.table,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   346
  functions: fun_spec History.T,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   347
  cases: case_spec History.T
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   348
};
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   349
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   350
fun types_of (Specs {types, ...}) = types;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   351
fun pending_eqns_of (Specs {pending_eqns, ...}) = pending_eqns;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   352
fun functions_of (Specs {functions, ...}) = functions;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   353
fun cases_of (Specs {cases, ...}) = cases;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   354
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   355
fun make_specs (types, ((pending_eqns, functions), cases)) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   356
  Specs {types = types, pending_eqns = pending_eqns,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   357
    functions = functions, cases = cases};
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   358
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   359
val empty_specs =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   360
  make_specs (Symtab.empty, ((Symtab.empty, Symtab.empty), Symtab.empty));
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   361
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   362
fun map_specs f (Specs {types = types, pending_eqns = pending_eqns,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   363
    functions = functions, cases = cases}) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   364
  make_specs (f (types, ((pending_eqns, functions), cases)));
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   365
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   366
fun merge_specs (Specs {types = types1, pending_eqns = _,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   367
    functions = functions1, cases = cases1},
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   368
  Specs {types = types2, pending_eqns = _,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   369
    functions = functions2, cases = cases2}) =
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   370
  let
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   371
    val types = History.join join_same_types (types1, types2);
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   372
    val all_types = map (snd o snd) (History.all types);
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   373
    fun check_abstype (c, fun_spec) = case associated_abstype fun_spec of
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   374
        NONE => true
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   375
      | SOME (tyco, abs) => (case History.lookup types tyco of
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   376
          NONE => false
70358
a55cfc8566fd proper quasi-total merge
haftmann
parents: 68773
diff changeset
   377
        | SOME (_, Constructors _) => false
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   378
        | SOME (_, Abstractor {abstractor = (abs', _), projection, more_abstract_functions, ...}) =>
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   379
            abs = abs' andalso (c = projection orelse member (op =) more_abstract_functions c));
66324
haftmann
parents: 66310
diff changeset
   380
    fun check_datatypes (_, case_spec) =
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   381
      let
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   382
        val (tycos, required_constructors) = associated_datatypes case_spec;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   383
        val allowed_constructors =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   384
          tycos
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   385
          |> maps (these o Option.map (concrete_constructors_of o snd) o History.lookup types)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   386
          |> map fst;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   387
      in subset (op =) (required_constructors, allowed_constructors) end;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   388
    val all_constructors =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   389
      maps (fst o constructors_of) all_types;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   390
    val functions = History.join fst (functions1, functions2)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   391
      |> fold (History.suppress o fst) all_constructors
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   392
      |> History.suppress_except check_abstype;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   393
    val cases = History.join fst (cases1, cases2)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   394
      |> History.suppress_except check_datatypes;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   395
  in make_specs (types, ((Symtab.empty, functions), cases)) end;
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   396
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   397
val map_types = map_specs o apfst;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   398
val map_pending_eqns = map_specs o apsnd o apfst o apfst;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   399
val map_functions = map_specs o apsnd o apfst o apsnd;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   400
val map_cases = map_specs o apsnd o apsnd;
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   401
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   402
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   403
(* data slots dependent on executable code *)
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   404
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   405
(*private copy avoids potential conflict of table exceptions*)
31971
8c1b845ed105 renamed functor TableFun to Table, and GraphFun to Graph;
wenzelm
parents: 31962
diff changeset
   406
structure Datatab = Table(type key = int val ord = int_ord);
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   407
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   408
local
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   409
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   410
type kind = {empty: Any.T};
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   411
43684
85388f5570c4 prefer Synchronized.var;
wenzelm
parents: 43639
diff changeset
   412
val kinds = Synchronized.var "Code_Data" (Datatab.empty: kind Datatab.table);
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   413
43684
85388f5570c4 prefer Synchronized.var;
wenzelm
parents: 43639
diff changeset
   414
fun invoke f k =
85388f5570c4 prefer Synchronized.var;
wenzelm
parents: 43639
diff changeset
   415
  (case Datatab.lookup (Synchronized.value kinds) k of
85388f5570c4 prefer Synchronized.var;
wenzelm
parents: 43639
diff changeset
   416
    SOME kind => f kind
85388f5570c4 prefer Synchronized.var;
wenzelm
parents: 43639
diff changeset
   417
  | NONE => raise Fail "Invalid code data identifier");
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   418
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   419
in
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   420
34173
458ced35abb8 reduced code generator cache to the baremost minimum
haftmann
parents: 33977
diff changeset
   421
fun declare_data empty =
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   422
  let
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   423
    val k = serial ();
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   424
    val kind = {empty = empty};
43684
85388f5570c4 prefer Synchronized.var;
wenzelm
parents: 43639
diff changeset
   425
    val _ = Synchronized.change kinds (Datatab.update (k, kind));
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   426
  in k end;
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   427
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   428
fun invoke_init k = invoke (fn kind => #empty kind) k;
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   429
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   430
end; (*local*)
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   431
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   432
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   433
(* global theory store *)
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   434
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   435
local
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   436
51368
2ea5c7c2d825 tuned signature -- prefer terminology of Scala and Axiom;
wenzelm
parents: 49971
diff changeset
   437
type data = Any.T Datatab.table;
72057
ce3f26b4e790 clarified -- avoid non-standard extend/merge;
wenzelm
parents: 70494
diff changeset
   438
78053
64f81d011a90 clarified signature: avoid convoluted operations;
wenzelm
parents: 77928
diff changeset
   439
fun make_dataref () =
64f81d011a90 clarified signature: avoid convoluted operations;
wenzelm
parents: 77928
diff changeset
   440
  Synchronized.var "code data" (NONE : (data * Context.theory_id) option);
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   441
34244
03f8dcab55f3 code cache without copy; tuned
haftmann
parents: 34184
diff changeset
   442
structure Code_Data = Theory_Data
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   443
(
72057
ce3f26b4e790 clarified -- avoid non-standard extend/merge;
wenzelm
parents: 70494
diff changeset
   444
  type T = specs * (string * (data * Context.theory_id) option Synchronized.var);
78053
64f81d011a90 clarified signature: avoid convoluted operations;
wenzelm
parents: 77928
diff changeset
   445
  val empty =
64f81d011a90 clarified signature: avoid convoluted operations;
wenzelm
parents: 77928
diff changeset
   446
    (empty_specs, (Context.theory_long_name (Context.the_global_context ()), make_dataref ()));
72057
ce3f26b4e790 clarified -- avoid non-standard extend/merge;
wenzelm
parents: 70494
diff changeset
   447
  fun merge ((specs1, dataref), (specs2, _)) =
ce3f26b4e790 clarified -- avoid non-standard extend/merge;
wenzelm
parents: 70494
diff changeset
   448
    (merge_specs (specs1, specs2), dataref);
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   449
);
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   450
72057
ce3f26b4e790 clarified -- avoid non-standard extend/merge;
wenzelm
parents: 70494
diff changeset
   451
fun init_dataref thy =
78053
64f81d011a90 clarified signature: avoid convoluted operations;
wenzelm
parents: 77928
diff changeset
   452
  let val thy_name = Context.theory_long_name thy in
64f81d011a90 clarified signature: avoid convoluted operations;
wenzelm
parents: 77928
diff changeset
   453
    if #1 (#2 (Code_Data.get thy)) = thy_name then NONE
64f81d011a90 clarified signature: avoid convoluted operations;
wenzelm
parents: 77928
diff changeset
   454
    else SOME ((Code_Data.map o apsnd) (K (thy_name, make_dataref ())) thy)
64f81d011a90 clarified signature: avoid convoluted operations;
wenzelm
parents: 77928
diff changeset
   455
  end;
72057
ce3f26b4e790 clarified -- avoid non-standard extend/merge;
wenzelm
parents: 70494
diff changeset
   456
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   457
in
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   458
72057
ce3f26b4e790 clarified -- avoid non-standard extend/merge;
wenzelm
parents: 70494
diff changeset
   459
val _ = Theory.setup (Theory.at_begin init_dataref);
ce3f26b4e790 clarified -- avoid non-standard extend/merge;
wenzelm
parents: 70494
diff changeset
   460
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   461
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   462
(* access to executable specifications *)
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   463
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   464
val specs_of : theory -> specs = fst o Code_Data.get;
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   465
72057
ce3f26b4e790 clarified -- avoid non-standard extend/merge;
wenzelm
parents: 70494
diff changeset
   466
fun modify_specs f thy =
78053
64f81d011a90 clarified signature: avoid convoluted operations;
wenzelm
parents: 77928
diff changeset
   467
  let val thy_name = Context.theory_long_name thy
64f81d011a90 clarified signature: avoid convoluted operations;
wenzelm
parents: 77928
diff changeset
   468
  in Code_Data.map (fn (specs, _) => (f specs, (thy_name, make_dataref ()))) thy end;
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   469
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   470
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   471
(* access to data dependent on executable specifications *)
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   472
34244
03f8dcab55f3 code cache without copy; tuned
haftmann
parents: 34184
diff changeset
   473
fun change_yield_data (kind, mk, dest) theory f =
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   474
  let
72057
ce3f26b4e790 clarified -- avoid non-standard extend/merge;
wenzelm
parents: 70494
diff changeset
   475
    val dataref = #2 (#2 (Code_Data.get theory));
67646
85582dded912 trim context of persistent data;
wenzelm
parents: 67522
diff changeset
   476
    val (datatab, thy_id) = case Synchronized.value dataref
85582dded912 trim context of persistent data;
wenzelm
parents: 67522
diff changeset
   477
     of SOME (datatab, thy_id) =>
85582dded912 trim context of persistent data;
wenzelm
parents: 67522
diff changeset
   478
        if Context.eq_thy_id (Context.theory_id theory, thy_id)
85582dded912 trim context of persistent data;
wenzelm
parents: 67522
diff changeset
   479
          then (datatab, thy_id)
85582dded912 trim context of persistent data;
wenzelm
parents: 67522
diff changeset
   480
          else (Datatab.empty, Context.theory_id theory)
85582dded912 trim context of persistent data;
wenzelm
parents: 67522
diff changeset
   481
      | NONE => (Datatab.empty, Context.theory_id theory)
34244
03f8dcab55f3 code cache without copy; tuned
haftmann
parents: 34184
diff changeset
   482
    val data = case Datatab.lookup datatab kind
03f8dcab55f3 code cache without copy; tuned
haftmann
parents: 34184
diff changeset
   483
     of SOME data => data
03f8dcab55f3 code cache without copy; tuned
haftmann
parents: 34184
diff changeset
   484
      | NONE => invoke_init kind;
40758
4f2c3e842ef8 consider sort constraints for datatype constructors when constructing the empty equation certificate;
haftmann
parents: 40726
diff changeset
   485
    val result as (_, data') = f (dest data);
34244
03f8dcab55f3 code cache without copy; tuned
haftmann
parents: 34184
diff changeset
   486
    val _ = Synchronized.change dataref
67646
85582dded912 trim context of persistent data;
wenzelm
parents: 67522
diff changeset
   487
      ((K o SOME) (Datatab.update (kind, mk data') datatab, thy_id));
34244
03f8dcab55f3 code cache without copy; tuned
haftmann
parents: 34184
diff changeset
   488
  in result end;
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   489
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   490
end; (*local*)
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   491
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   492
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   493
(* pending function equations *)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   494
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   495
(* Ideally, *all* equations implementing a functions would be treated as
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   496
   *one* atomic declaration;  unfortunately, we cannot implement this:
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   497
   the too-well-established declaration interface are Isar attributes
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   498
   which operate on *one* single theorem.  Hence we treat such Isar
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   499
   declarations as "pending" and historize them as proper declarations
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   500
   at the end of each theory. *)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   501
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   502
fun modify_pending_eqns thy { check_singleton } c f =
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   503
  map_pending_eqns (Symtab.map_default (c, []) (fn eqns =>
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   504
    if null eqns orelse not check_singleton orelse not (Config.get_global thy only_single_equation)
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   505
    then f eqns
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   506
    else error ("Only a single code equation is allowed for " ^ string_of_const thy c)
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   507
  ));
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   508
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   509
fun register_fun_spec c spec =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   510
  map_pending_eqns (Symtab.delete_safe c)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   511
  #> map_functions (History.register c spec);
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   512
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   513
fun lookup_fun_spec specs c =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   514
  case Symtab.lookup (pending_eqns_of specs) c of
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   515
    SOME eqns => Eqns (false, eqns)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   516
  | NONE => (case History.lookup (functions_of specs) c of
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   517
      SOME spec => spec
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   518
    | NONE => unimplemented);
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   519
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   520
fun lookup_proper_fun_spec specs c =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   521
  let
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   522
    val spec = lookup_fun_spec specs c
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   523
  in
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   524
    if is_unimplemented spec then NONE else SOME spec
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   525
  end;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   526
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   527
fun all_fun_specs specs =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   528
  map_filter (fn c => Option.map (pair c) (lookup_proper_fun_spec specs c))
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   529
    (union (op =)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   530
      ((Symtab.keys o pending_eqns_of) specs)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   531
      ((Symtab.keys o functions_of) specs));
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   532
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   533
fun historize_pending_fun_specs thy =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   534
  let
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   535
    val pending_eqns = (pending_eqns_of o specs_of) thy;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   536
  in if Symtab.is_empty pending_eqns
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   537
    then
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   538
      NONE
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   539
    else
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   540
      thy
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   541
      |> modify_specs (map_functions
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   542
          (Symtab.fold (fn (c, eqs) => History.register c (Eqns (false, eqs))) pending_eqns)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   543
          #> map_pending_eqns (K Symtab.empty))
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   544
      |> SOME
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   545
  end;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   546
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
   547
val _ = Theory.setup (Theory.at_end (historize_pending_fun_specs));
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   548
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   549
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   550
(** foundation **)
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   551
66128
0181bb24bdca more uniform ordering and naming of sections;
haftmann
parents: 66127
diff changeset
   552
(* types *)
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   553
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   554
fun no_constr thy s (c, ty) = error ("Not a datatype constructor:\n" ^ string_of_const thy c
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   555
  ^ " :: " ^ string_of_typ thy ty ^ "\n" ^ enclose "(" ")" s);
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   556
45987
9ba44b49859b dropped disfruitful `constant signatures`
haftmann
parents: 45430
diff changeset
   557
fun analyze_constructor thy (c, ty) =
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   558
  let
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59458
diff changeset
   559
    val _ = Thm.global_cterm_of thy (Const (c, ty));
54603
89d5b69e5a08 prefer name-normalizing devarify over unvarifyT whenever appropriate
haftmann
parents: 53171
diff changeset
   560
    val ty_decl = devarify (const_typ thy c);
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   561
    fun last_typ c_ty ty =
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   562
      let
33531
2c0a4adbcf13 tuned error messages; tuned code
haftmann
parents: 33522
diff changeset
   563
        val tfrees = Term.add_tfreesT ty [];
40844
5895c525739d more direct use of binder_types/body_type;
wenzelm
parents: 40803
diff changeset
   564
        val (tyco, vs) = (apsnd o map) dest_TFree (dest_Type (body_type ty))
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   565
          handle TYPE _ => no_constr thy "bad type" c_ty
36112
7fa17a225852 user interface for abstract datatypes is an attribute, not a command
haftmann
parents: 35845
diff changeset
   566
        val _ = if tyco = "fun" then no_constr thy "bad type" c_ty else ();
45344
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 45211
diff changeset
   567
        val _ =
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 45211
diff changeset
   568
          if has_duplicates (eq_fst (op =)) vs
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   569
          then no_constr thy "duplicate type variables in datatype" c_ty else ();
45344
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 45211
diff changeset
   570
        val _ =
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 45211
diff changeset
   571
          if length tfrees <> length vs
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   572
          then no_constr thy "type variables missing in datatype" c_ty else ();
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   573
      in (tyco, vs) end;
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   574
    val (tyco, _) = last_typ (c, ty) ty_decl;
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   575
    val (_, vs) = last_typ (c, ty) ty;
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   576
  in ((tyco, map snd vs), (c, (map fst vs, ty))) end;
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   577
49904
2df2786ac7b7 no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents: 49760
diff changeset
   578
fun constrset_of_consts thy consts =
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   579
  let
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51584
diff changeset
   580
    val _ = map (fn (c, _) => if (is_some o Axclass.class_of_param thy) c
49904
2df2786ac7b7 no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents: 49760
diff changeset
   581
      then error ("Is a class parameter: " ^ string_of_const thy c) else ()) consts;
2df2786ac7b7 no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents: 49760
diff changeset
   582
    val raw_constructors = map (analyze_constructor thy) consts;
2df2786ac7b7 no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents: 49760
diff changeset
   583
    val tyco = case distinct (op =) (map (fst o fst) raw_constructors)
2df2786ac7b7 no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents: 49760
diff changeset
   584
     of [tyco] => tyco
2df2786ac7b7 no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents: 49760
diff changeset
   585
      | [] => error "Empty constructor set"
2df2786ac7b7 no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents: 49760
diff changeset
   586
      | tycos => error ("Different type constructors in constructor set: " ^ commas_quote tycos)
81507
08574da77b4a clarified signature: shorten common cases;
wenzelm
parents: 79411
diff changeset
   587
    val vs = Name.invent_global_types (Sign.arity_number thy tyco);
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   588
    fun inst vs' (c, (vs, ty)) =
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   589
      let
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   590
        val the_v = the o AList.lookup (op =) (vs ~~ vs');
49904
2df2786ac7b7 no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents: 49760
diff changeset
   591
        val ty' = map_type_tfree (fn (v, _) => TFree (the_v v, [])) ty;
2df2786ac7b7 no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents: 49760
diff changeset
   592
        val (vs'', ty'') = typscheme thy (c, ty');
2df2786ac7b7 no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents: 49760
diff changeset
   593
      in (c, (vs'', binder_types ty'')) end;
2df2786ac7b7 no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents: 49760
diff changeset
   594
    val constructors = map (inst vs o snd) raw_constructors;
2df2786ac7b7 no sort constraints on datatype constructors in internal bookkeeping
haftmann
parents: 49760
diff changeset
   595
  in (tyco, (map (rpair []) vs, constructors)) end;
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   596
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   597
fun lookup_vs_type_spec thy = History.lookup ((types_of o specs_of) thy);
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   598
66328
cf9ce8016da1 clarified
haftmann
parents: 66324
diff changeset
   599
type constructors =
cf9ce8016da1 clarified
haftmann
parents: 66324
diff changeset
   600
  (string * sort) list * (string * ((string * sort) list * typ list)) list;
cf9ce8016da1 clarified
haftmann
parents: 66324
diff changeset
   601
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   602
fun get_type thy tyco = case lookup_vs_type_spec thy tyco
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   603
 of SOME (vs, type_spec) => apfst (pair vs) (constructors_of type_spec)
45987
9ba44b49859b dropped disfruitful `constant signatures`
haftmann
parents: 45430
diff changeset
   604
  | NONE => Sign.arity_number thy tyco
81507
08574da77b4a clarified signature: shorten common cases;
wenzelm
parents: 79411
diff changeset
   605
      |> Name.invent_global_types
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   606
      |> map (rpair [])
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   607
      |> rpair []
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   608
      |> rpair false;
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   609
66328
cf9ce8016da1 clarified
haftmann
parents: 66324
diff changeset
   610
type abs_type =
cf9ce8016da1 clarified
haftmann
parents: 66324
diff changeset
   611
  (string * sort) list * {abs_rep: thm, abstractor: string * ((string * sort) list * typ), projection: string};
cf9ce8016da1 clarified
haftmann
parents: 66324
diff changeset
   612
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   613
fun get_abstype_spec thy tyco = case lookup_vs_type_spec thy tyco of
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   614
    SOME (vs, Abstractor {abs_rep, abstractor, projection, ...}) =>
78056
904242f46ce1 more accurate Thm.trim_context / Thm.transfer;
wenzelm
parents: 78054
diff changeset
   615
      (vs, {abs_rep = Thm.transfer thy abs_rep, abstractor = abstractor, projection = projection})
36122
45f8898fe4cf more accurate pattern match
haftmann
parents: 36112
diff changeset
   616
  | _ => error ("Not an abstract type: " ^ tyco);
66167
1bd268ab885c more informative task_statistics;
wenzelm
parents: 66149
diff changeset
   617
35299
4f4d5bf4ea08 proper distinction of code datatypes and abstypes
haftmann
parents: 35226
diff changeset
   618
fun get_type_of_constr_or_abstr thy c =
40844
5895c525739d more direct use of binder_types/body_type;
wenzelm
parents: 40803
diff changeset
   619
  case (body_type o const_typ thy) c
40758
4f2c3e842ef8 consider sort constraints for datatype constructors when constructing the empty equation certificate;
haftmann
parents: 40726
diff changeset
   620
   of Type (tyco, _) => let val ((_, cos), abstract) = get_type thy tyco
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   621
        in if member (op =) (map fst cos) c then SOME (tyco, abstract) else NONE end
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   622
    | _ => NONE;
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   623
35299
4f4d5bf4ea08 proper distinction of code datatypes and abstypes
haftmann
parents: 35226
diff changeset
   624
fun is_constr thy c = case get_type_of_constr_or_abstr thy c
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   625
 of SOME (_, false) => true
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   626
   | _ => false;
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   627
35299
4f4d5bf4ea08 proper distinction of code datatypes and abstypes
haftmann
parents: 35226
diff changeset
   628
fun is_abstr thy c = case get_type_of_constr_or_abstr thy c
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   629
 of SOME (_, true) => true
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   630
   | _ => false;
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   631
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   632
34874
89c230bf8feb added code certificates
haftmann
parents: 34272
diff changeset
   633
(* bare code equations *)
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   634
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   635
(* convention for variables:
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   636
    ?x ?'a   for free-floating theorems (e.g. in the data store)
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   637
    ?x  'a   for certificates
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   638
     x  'a   for final representation of equations
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   639
*)
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   640
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   641
exception BAD_THM of string;
63242
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   642
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   643
datatype strictness = Silent | Liberal | Strict
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   644
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   645
fun handle_strictness thm_of f strictness thy x = SOME (f x)
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   646
  handle BAD_THM msg => case strictness of
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   647
    Silent => NONE
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   648
  | Liberal => (warning (msg ^ ", in theorem:\n" ^ Thm.string_of_thm_global thy (thm_of x)); NONE)
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   649
  | Strict => error (msg ^ ", in theorem:\n" ^ Thm.string_of_thm_global thy (thm_of x));
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   650
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   651
fun is_linear thm =
63242
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   652
  let
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   653
    val (_, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   654
  in
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   655
    not (has_duplicates (op =) ((fold o fold_aterms)
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   656
      (fn Var (v, _) => cons v | _ => I) args []))
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   657
  end;
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   658
36209
566be5d48090 more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
haftmann
parents: 36202
diff changeset
   659
fun check_decl_ty thy (c, ty) =
566be5d48090 more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
haftmann
parents: 36202
diff changeset
   660
  let
54604
1512fa5fe531 prefer sort-stripping const_typ over Sign.the_const_type whenever appropriate
haftmann
parents: 54603
diff changeset
   661
    val ty_decl = const_typ thy c;
49534
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   662
  in if typscheme_equiv (ty_decl, ty) then ()
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   663
    else raise BAD_THM ("Type\n" ^ string_of_typ thy ty
36209
566be5d48090 more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
haftmann
parents: 36202
diff changeset
   664
      ^ "\nof constant " ^ quote c
40362
82a066bff182 Code.check_const etc.: reject too specific types
haftmann
parents: 40316
diff changeset
   665
      ^ "\nis too specific compared to declared type\n"
36209
566be5d48090 more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
haftmann
parents: 36202
diff changeset
   666
      ^ string_of_typ thy ty_decl)
66167
1bd268ab885c more informative task_statistics;
wenzelm
parents: 66149
diff changeset
   667
  end;
36209
566be5d48090 more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
haftmann
parents: 36202
diff changeset
   668
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   669
fun check_eqn thy {allow_nonlinear, allow_consts, allow_pats} thm (lhs, rhs) =
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   670
  let
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   671
    fun vars_of t = fold_aterms (fn Var (v, _) => insert (op =) v
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   672
      | Free _ => raise BAD_THM "Illegal free variable"
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   673
      | _ => I) t [];
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   674
    fun tvars_of t = fold_term_types (fn _ =>
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   675
      fold_atyps (fn TVar (v, _) => insert (op =) v
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   676
        | TFree _ => raise BAD_THM "Illegal free type variable")) t [];
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   677
    val lhs_vs = vars_of lhs;
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   678
    val rhs_vs = vars_of rhs;
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   679
    val lhs_tvs = tvars_of lhs;
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   680
    val rhs_tvs = tvars_of rhs;
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   681
    val _ = if null (subtract (op =) lhs_vs rhs_vs)
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   682
      then ()
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   683
      else raise BAD_THM "Free variables on right hand side of equation";
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   684
    val _ = if null (subtract (op =) lhs_tvs rhs_tvs)
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   685
      then ()
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   686
      else raise BAD_THM "Free type variables on right hand side of equation";
34894
fadbdd350dd1 corrected error messages; tuned
haftmann
parents: 34891
diff changeset
   687
    val (head, args) = strip_comb lhs;
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   688
    val (c, ty) = case head
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51584
diff changeset
   689
     of Const (c_ty as (_, ty)) => (Axclass.unoverload_const thy c_ty, ty)
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   690
      | _ => raise BAD_THM "Equation not headed by constant";
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   691
    fun check _ (Abs _) = raise BAD_THM "Abstraction on left hand side of equation"
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   692
      | check 0 (Var _) = ()
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   693
      | check _ (Var _) = raise BAD_THM "Variable with application on left hand side of equation"
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   694
      | check n (t1 $ t2) = (check (n+1) t1; check 0 t2)
34894
fadbdd350dd1 corrected error messages; tuned
haftmann
parents: 34891
diff changeset
   695
      | check n (Const (c_ty as (c, ty))) =
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   696
          if allow_pats then let
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51584
diff changeset
   697
            val c' = Axclass.unoverload_const thy c_ty
45987
9ba44b49859b dropped disfruitful `constant signatures`
haftmann
parents: 45430
diff changeset
   698
          in if n = (length o binder_types) ty
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   699
            then if allow_consts orelse is_constr thy c'
33940
317933ce3712 crude support for type aliasses and corresponding constant signatures
haftmann
parents: 33756
diff changeset
   700
              then ()
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   701
              else raise BAD_THM (quote c ^ " is not a constructor, on left hand side of equation")
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   702
            else raise BAD_THM ("Partially applied constant " ^ quote c ^ " on left hand side of equation")
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   703
          end else raise BAD_THM ("Pattern not allowed here, but constant " ^ quote c ^ " encountered on left hand side of equation")
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   704
    val _ = map (check 0) args;
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   705
    val _ = if allow_nonlinear orelse is_linear thm then ()
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   706
      else raise BAD_THM "Duplicate variables on left hand side of equation";
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51584
diff changeset
   707
    val _ = if (is_none o Axclass.class_of_param thy) c then ()
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   708
      else raise BAD_THM "Overloaded constant as head in equation";
34894
fadbdd350dd1 corrected error messages; tuned
haftmann
parents: 34891
diff changeset
   709
    val _ = if not (is_constr thy c) then ()
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   710
      else raise BAD_THM "Constructor as head in equation";
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   711
    val _ = if not (is_abstr thy c) then ()
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   712
      else raise BAD_THM "Abstractor as head in equation";
36209
566be5d48090 more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
haftmann
parents: 36202
diff changeset
   713
    val _ = check_decl_ty thy (c, ty);
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   714
    val _ = case strip_type ty of
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   715
        (Type (tyco, _) :: _, _) => (case lookup_vs_type_spec thy tyco of
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   716
          SOME (_, type_spec) => (case projection_of type_spec of
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   717
            SOME proj =>
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   718
              if c = proj
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   719
              then raise BAD_THM "Projection as head in equation"
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   720
              else ()
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   721
          | _ => ())
52475
445ae7a4e4e1 sort out code equations headed by a projection of a abstract datatype
haftmann
parents: 51717
diff changeset
   722
        | _ => ())
445ae7a4e4e1 sort out code equations headed by a projection of a abstract datatype
haftmann
parents: 51717
diff changeset
   723
      | _ => ();
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   724
  in () end;
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   725
63242
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   726
local
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   727
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   728
fun raw_assert_eqn thy check_patterns (thm, proper) =
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   729
  let
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   730
    val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   731
      handle TERM _ => raise BAD_THM "Not an equation"
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   732
           | THM _ => raise BAD_THM "Not a proper equation";
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   733
    val _ = check_eqn thy {allow_nonlinear = not proper,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   734
      allow_consts = not (proper andalso check_patterns), allow_pats = true} thm (lhs, rhs);
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   735
  in (thm, proper) end;
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   736
63242
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   737
fun raw_assert_abs_eqn thy some_tyco thm =
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   738
  let
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   739
    val (full_lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   740
      handle TERM _ => raise BAD_THM "Not an equation"
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   741
           | THM _ => raise BAD_THM "Not a proper equation";
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   742
    val (proj_t, lhs) = dest_comb full_lhs
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   743
      handle TERM _ => raise BAD_THM "Not an abstract equation";
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   744
    val (proj, ty) = dest_Const proj_t
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   745
      handle TERM _ => raise BAD_THM "Not an abstract equation";
40187
9b6e918682d5 more general treatment of type argument in code certificates for operations on abstract types
haftmann
parents: 39811
diff changeset
   746
    val (tyco, Ts) = (dest_Type o domain_type) ty
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   747
      handle TERM _ => raise BAD_THM "Not an abstract equation"
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   748
           | TYPE _ => raise BAD_THM "Not an abstract equation";
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   749
    val _ = case some_tyco of SOME tyco' => if tyco = tyco' then ()
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   750
          else raise BAD_THM ("Abstract type mismatch:" ^ quote tyco ^ " vs. " ^ quote tyco')
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   751
      | NONE => ();
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   752
    val (vs, proj', (abs', _)) = case lookup_vs_type_spec thy tyco
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   753
     of SOME (vs, Abstractor spec) => (vs, #projection spec, #abstractor spec)
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   754
      | _ => raise BAD_THM ("Not an abstract type: " ^ tyco);
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   755
    val _ = if proj = proj' then ()
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   756
      else raise BAD_THM ("Projection mismatch: " ^ quote proj ^ " vs. " ^ quote proj');
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   757
    val _ = check_eqn thy {allow_nonlinear = false,
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   758
      allow_consts = false, allow_pats = false} thm (lhs, rhs);
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   759
    val _ = if ListPair.all (fn (T, (_, sort)) => Sign.of_sort thy (T, sort)) (Ts, vs) then ()
40187
9b6e918682d5 more general treatment of type argument in code certificates for operations on abstract types
haftmann
parents: 39811
diff changeset
   760
      else error ("Type arguments do not satisfy sort constraints of abstype certificate.");
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   761
  in (thm, (tyco, abs')) end;
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   762
63242
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   763
in
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   764
66260
haftmann
parents: 66251
diff changeset
   765
fun generic_assert_eqn strictness thy check_patterns eqn =
haftmann
parents: 66251
diff changeset
   766
  handle_strictness fst (raw_assert_eqn thy check_patterns) strictness thy eqn;
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   767
63242
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   768
fun generic_assert_abs_eqn strictness thy check_patterns thm =
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   769
  handle_strictness I (raw_assert_abs_eqn thy check_patterns) strictness thy thm;
52637
1501ebe39711 attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents: 52475
diff changeset
   770
63242
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   771
end;
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   772
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   773
fun assert_eqn thy = the o generic_assert_eqn Strict thy true;
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   774
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   775
fun assert_abs_eqn thy some_tyco = the o generic_assert_abs_eqn Strict thy some_tyco;
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   776
33940
317933ce3712 crude support for type aliasses and corresponding constant signatures
haftmann
parents: 33756
diff changeset
   777
val head_eqn = dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of;
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   778
31957
a9742afd403e tuned interface of structure Code
haftmann
parents: 31890
diff changeset
   779
fun const_typ_eqn thy thm =
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   780
  let
32640
ba6531df2c64 corrected order of type variables in code equations; more precise certificate for cases
haftmann
parents: 32544
diff changeset
   781
    val (c, ty) = head_eqn thm;
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51584
diff changeset
   782
    val c' = Axclass.unoverload_const thy (c, ty);
33940
317933ce3712 crude support for type aliasses and corresponding constant signatures
haftmann
parents: 33756
diff changeset
   783
      (*permissive wrt. to overloaded constants!*)
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   784
  in (c', ty) end;
33940
317933ce3712 crude support for type aliasses and corresponding constant signatures
haftmann
parents: 33756
diff changeset
   785
31957
a9742afd403e tuned interface of structure Code
haftmann
parents: 31890
diff changeset
   786
fun const_eqn thy = fst o const_typ_eqn thy;
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   787
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51584
diff changeset
   788
fun const_abs_eqn thy = Axclass.unoverload_const thy o dest_Const o fst o strip_comb o snd
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   789
  o dest_comb o fst o Logic.dest_equals o Thm.plain_prop_of;
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   790
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   791
fun mk_proj tyco vs ty abs rep =
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   792
  let
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   793
    val ty_abs = Type (tyco, map TFree vs);
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   794
    val xarg = Var (("x", 0), ty);
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   795
  in Logic.mk_equals (Const (rep, ty_abs --> ty) $ (Const (abs, ty --> ty_abs) $ xarg), xarg) end;
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   796
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   797
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   798
(* technical transformations of code equations *)
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   799
63242
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   800
fun meta_rewrite thy = Local_Defs.meta_rewrite_rule (Proof_Context.init_global thy);
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   801
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   802
fun same_arity thy thms =
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
   803
  let
77928
faaff590bd9e stripped unused functionality
haftmann
parents: 77702
diff changeset
   804
    val lhs_rhss = map (Logic.dest_equals o Thm.plain_prop_of) thms;
faaff590bd9e stripped unused functionality
haftmann
parents: 77702
diff changeset
   805
    val k = fold (Integer.max o length o snd o strip_comb o fst) lhs_rhss 0;
faaff590bd9e stripped unused functionality
haftmann
parents: 77702
diff changeset
   806
    fun expand_eta (lhs, rhs) thm =
faaff590bd9e stripped unused functionality
haftmann
parents: 77702
diff changeset
   807
      let
faaff590bd9e stripped unused functionality
haftmann
parents: 77702
diff changeset
   808
        val l = k - length (snd (strip_comb lhs));
faaff590bd9e stripped unused functionality
haftmann
parents: 77702
diff changeset
   809
        val (raw_vars, _) = Term.strip_abs_eta l rhs;
faaff590bd9e stripped unused functionality
haftmann
parents: 77702
diff changeset
   810
        val vars = burrow_fst (Name.variant_list (map (fst o fst) (Term.add_vars lhs [])))
faaff590bd9e stripped unused functionality
haftmann
parents: 77702
diff changeset
   811
          raw_vars;
faaff590bd9e stripped unused functionality
haftmann
parents: 77702
diff changeset
   812
        fun expand (v, ty) thm = Drule.fun_cong_rule thm
faaff590bd9e stripped unused functionality
haftmann
parents: 77702
diff changeset
   813
          (Thm.global_cterm_of thy (Var ((v, 0), ty)));
faaff590bd9e stripped unused functionality
haftmann
parents: 77702
diff changeset
   814
      in
faaff590bd9e stripped unused functionality
haftmann
parents: 77702
diff changeset
   815
        thm
faaff590bd9e stripped unused functionality
haftmann
parents: 77702
diff changeset
   816
        |> fold expand vars
faaff590bd9e stripped unused functionality
haftmann
parents: 77702
diff changeset
   817
        |> Conv.fconv_rule Drule.beta_eta_conversion
faaff590bd9e stripped unused functionality
haftmann
parents: 77702
diff changeset
   818
      end;
faaff590bd9e stripped unused functionality
haftmann
parents: 77702
diff changeset
   819
  in map2 expand_eta lhs_rhss thms end;
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   820
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   821
fun mk_desymbolization pre post mk vs =
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   822
  let
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   823
    val names = map (pre o fst o fst) vs
56811
b66639331db5 optional case enforcement
haftmann
parents: 56375
diff changeset
   824
      |> map (Name.desymbolize (SOME false))
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   825
      |> Name.variant_list []
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   826
      |> map post;
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   827
  in map_filter (fn (((v, i), x), v') =>
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   828
    if v = v' andalso i = 0 then NONE
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   829
    else SOME (((v, i), x), mk ((v', 0), x))) (vs ~~ names)
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   830
  end;
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   831
60367
e201bd8973db clarified context;
wenzelm
parents: 59787
diff changeset
   832
fun desymbolize_tvars thy thms =
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   833
  let
75353
05f7f5454cb6 prefer build combinator
haftmann
parents: 74978
diff changeset
   834
    val tvs = build (fold (Term.add_tvars o Thm.prop_of) thms);
60805
4cc49ead6e75 tuned signature;
wenzelm
parents: 60367
diff changeset
   835
    val instT =
4cc49ead6e75 tuned signature;
wenzelm
parents: 60367
diff changeset
   836
      mk_desymbolization (unprefix "'") (prefix "'") (Thm.global_ctyp_of thy o TVar) tvs;
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74235
diff changeset
   837
  in map (Thm.instantiate (TVars.make instT, Vars.empty)) thms end;
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   838
60367
e201bd8973db clarified context;
wenzelm
parents: 59787
diff changeset
   839
fun desymbolize_vars thy thm =
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   840
  let
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   841
    val vs = Term.add_vars (Thm.prop_of thm) [];
60805
4cc49ead6e75 tuned signature;
wenzelm
parents: 60367
diff changeset
   842
    val inst = mk_desymbolization I I (Thm.global_cterm_of thy o Var) vs;
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74235
diff changeset
   843
  in Thm.instantiate (TVars.empty, Vars.make inst) thm end;
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   844
60367
e201bd8973db clarified context;
wenzelm
parents: 59787
diff changeset
   845
fun canonize_thms thy = desymbolize_tvars thy #> same_arity thy #> map (desymbolize_vars thy);
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
   846
34874
89c230bf8feb added code certificates
haftmann
parents: 34272
diff changeset
   847
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   848
(* preparation and classification of code equations *)
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   849
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   850
fun prep_eqn strictness thy =
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   851
  apfst (meta_rewrite thy)
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   852
  #> generic_assert_eqn strictness thy false
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   853
  #> Option.map (fn eqn => (const_eqn thy (fst eqn), eqn));
67032
ed499d1252fc strip some trailing spaces to force Pure rebuild after ce6454669360
Lars Hupel <lars.hupel@mytum.de>
parents: 66332
diff changeset
   854
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   855
fun prep_eqns strictness thy =
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   856
  map_filter (prep_eqn strictness thy)
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   857
  #> AList.group (op =);
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   858
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   859
fun prep_abs_eqn strictness thy =
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   860
  meta_rewrite thy
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   861
  #> generic_assert_abs_eqn strictness thy NONE
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   862
  #> Option.map (fn abs_eqn => (const_abs_eqn thy (fst abs_eqn), abs_eqn));
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   863
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   864
fun prep_maybe_abs_eqn thy raw_thm =
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   865
  let
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   866
    val thm = meta_rewrite thy raw_thm;
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   867
    val some_abs_thm = generic_assert_abs_eqn Silent thy NONE thm;
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   868
  in case some_abs_thm of
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   869
      SOME (thm, tyco) => SOME (const_abs_eqn thy thm, ((thm, true), SOME tyco))
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   870
    | NONE => generic_assert_eqn Liberal thy false (thm, false)
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   871
        |> Option.map (fn (thm, _) => (const_eqn thy thm, ((thm, is_linear thm), NONE)))
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   872
  end;
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   873
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
   874
36112
7fa17a225852 user interface for abstract datatypes is an attribute, not a command
haftmann
parents: 35845
diff changeset
   875
(* abstype certificates *)
7fa17a225852 user interface for abstract datatypes is an attribute, not a command
haftmann
parents: 35845
diff changeset
   876
63242
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   877
local
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   878
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   879
fun raw_abstype_cert thy proto_thm =
36112
7fa17a225852 user interface for abstract datatypes is an attribute, not a command
haftmann
parents: 35845
diff changeset
   880
  let
63073
413184c7a2a2 clarified context, notably for internal use of Simplifier;
wenzelm
parents: 61268
diff changeset
   881
    val thm = (Axclass.unoverload (Proof_Context.init_global thy) o meta_rewrite thy) proto_thm;
36112
7fa17a225852 user interface for abstract datatypes is an attribute, not a command
haftmann
parents: 35845
diff changeset
   882
    val (lhs, rhs) = Logic.dest_equals (Thm.plain_prop_of thm)
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   883
      handle TERM _ => raise BAD_THM "Not an equation"
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   884
           | THM _ => raise BAD_THM "Not a proper equation";
36209
566be5d48090 more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
haftmann
parents: 36202
diff changeset
   885
    val ((abs, raw_ty), ((rep, rep_ty), param)) = (apsnd (apfst dest_Const o dest_comb)
36112
7fa17a225852 user interface for abstract datatypes is an attribute, not a command
haftmann
parents: 35845
diff changeset
   886
        o apfst dest_Const o dest_comb) lhs
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   887
      handle TERM _ => raise BAD_THM "Not an abstype certificate";
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58666
diff changeset
   888
    val _ = apply2 (fn c => if (is_some o Axclass.class_of_param thy) c
36209
566be5d48090 more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
haftmann
parents: 36202
diff changeset
   889
      then error ("Is a class parameter: " ^ string_of_const thy c) else ()) (abs, rep);
566be5d48090 more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
haftmann
parents: 36202
diff changeset
   890
    val _ = check_decl_ty thy (abs, raw_ty);
566be5d48090 more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
haftmann
parents: 36202
diff changeset
   891
    val _ = check_decl_ty thy (rep, rep_ty);
48068
04aeda922be2 explicit check for correct number of arguments for abstract constructor
haftmann
parents: 47555
diff changeset
   892
    val _ = if length (binder_types raw_ty) = 1
04aeda922be2 explicit check for correct number of arguments for abstract constructor
haftmann
parents: 47555
diff changeset
   893
      then ()
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   894
      else raise BAD_THM "Bad type for abstract constructor";
40758
4f2c3e842ef8 consider sort constraints for datatype constructors when constructing the empty equation certificate;
haftmann
parents: 40726
diff changeset
   895
    val _ = (fst o dest_Var) param
78139
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   896
      handle TERM _ => raise BAD_THM "Not an abstype certificate";
bb85bda12b8e proper exception positions;
wenzelm
parents: 78095
diff changeset
   897
    val _ = if param = rhs then () else raise BAD_THM "Not an abstype certificate";
45344
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 45211
diff changeset
   898
    val ((tyco, sorts), (abs, (vs, ty'))) =
54603
89d5b69e5a08 prefer name-normalizing devarify over unvarifyT whenever appropriate
haftmann
parents: 53171
diff changeset
   899
      analyze_constructor thy (abs, devarify raw_ty);
36112
7fa17a225852 user interface for abstract datatypes is an attribute, not a command
haftmann
parents: 35845
diff changeset
   900
    val ty = domain_type ty';
49534
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   901
    val (vs', _) = typscheme thy (abs, ty');
40726
16dcfedc4eb7 keep type variable arguments of datatype constructors in bookkeeping
haftmann
parents: 40564
diff changeset
   902
  in (tyco, (vs ~~ sorts, ((abs, (vs', ty)), (rep, thm)))) end;
36112
7fa17a225852 user interface for abstract datatypes is an attribute, not a command
haftmann
parents: 35845
diff changeset
   903
63242
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   904
in
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   905
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   906
fun check_abstype_cert strictness thy proto_thm =
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   907
  handle_strictness I (raw_abstype_cert thy) strictness thy proto_thm;
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   908
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   909
end;
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
   910
36112
7fa17a225852 user interface for abstract datatypes is an attribute, not a command
haftmann
parents: 35845
diff changeset
   911
34874
89c230bf8feb added code certificates
haftmann
parents: 34272
diff changeset
   912
(* code equation certificates *)
89c230bf8feb added code certificates
haftmann
parents: 34272
diff changeset
   913
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   914
fun build_head thy (c, ty) =
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59458
diff changeset
   915
  Thm.global_cterm_of thy (Logic.mk_equals (Free ("HEAD", ty), Const (c, ty)));
34874
89c230bf8feb added code certificates
haftmann
parents: 34272
diff changeset
   916
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   917
fun get_head thy cert_thm =
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   918
  let
60949
ccbf9379e355 added Thm.chyps_of;
wenzelm
parents: 60948
diff changeset
   919
    val [head] = Thm.chyps_of cert_thm;
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   920
    val (_, Const (c, ty)) = (Logic.dest_equals o Thm.term_of) head;
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   921
  in (typscheme thy (c, ty), head) end;
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   922
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   923
fun typscheme_projection thy =
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   924
  typscheme thy o dest_Const o fst o dest_comb o fst o Logic.dest_equals;
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   925
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   926
fun typscheme_abs thy =
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   927
  typscheme thy o dest_Const o fst o strip_comb o snd o dest_comb o fst o Logic.dest_equals o Thm.prop_of;
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   928
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   929
fun constrain_thm thy vs sorts thm =
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   930
  let
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   931
    val mapping = map2 (fn (v, sort) => fn sort' =>
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   932
      (v, Sorts.inter_sort (Sign.classes_of thy) (sort, sort'))) vs sorts;
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74235
diff changeset
   933
    val instT =
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74235
diff changeset
   934
      TVars.build
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74235
diff changeset
   935
        (fold2 (fn (v, sort) => fn (_, sort') =>
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74235
diff changeset
   936
          TVars.add (((v, 0), sort), Thm.global_ctyp_of thy (TFree (v, sort')))) vs mapping);
66130
d0476618a94c more consistent terminology
haftmann
parents: 66129
diff changeset
   937
    val subst = (Term.map_types o map_type_tfree)
40803
haftmann
parents: 40764
diff changeset
   938
      (fn (v, _) => TFree (v, the (AList.lookup (op =) mapping v)));
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   939
  in
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   940
    thm
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35624
diff changeset
   941
    |> Thm.varifyT_global
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74235
diff changeset
   942
    |> Thm.instantiate (instT, Vars.empty)
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   943
    |> pair subst
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   944
  end;
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   945
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   946
fun concretify_abs thy tyco abs_thm =
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   947
  let
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   948
    val (_, {abstractor = (c_abs, _), abs_rep, ...}) = get_abstype_spec thy tyco;
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   949
    val lhs = (fst o Logic.dest_equals o Thm.prop_of) abs_thm
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   950
    val ty = fastype_of lhs;
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   951
    val ty_abs = (fastype_of o snd o dest_comb) lhs;
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   952
    val abs = Thm.global_cterm_of thy (Const (c_abs, ty --> ty_abs));
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   953
    val raw_concrete_thm = Drule.transitive_thm OF [Thm.symmetric abs_rep, Thm.combination (Thm.reflexive abs) abs_thm];
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
   954
  in (c_abs, (Thm.varifyT_global o zero_var_indexes) raw_concrete_thm) end;
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   955
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   956
fun add_rhss_of_eqn thy t =
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   957
  let
45987
9ba44b49859b dropped disfruitful `constant signatures`
haftmann
parents: 45430
diff changeset
   958
    val (args, rhs) = (apfst (snd o strip_comb) o Logic.dest_equals) t;
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   959
    fun add_const (Const (c, ty)) = insert (op =) (c, Sign.const_typargs thy (c, ty))
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   960
      | add_const _ = I
39568
839a0446630b corrected long-overlooked slip: the Pure equality of a code equation is no part of the code equation itself
haftmann
parents: 39552
diff changeset
   961
    val add_consts = fold_aterms add_const
839a0446630b corrected long-overlooked slip: the Pure equality of a code equation is no part of the code equation itself
haftmann
parents: 39552
diff changeset
   962
  in add_consts rhs o fold add_consts args end;
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   963
46513
2659ee0128c2 more explicit error on malformed abstract equation; dropped dead code; tuned signature
haftmann
parents: 45987
diff changeset
   964
val dest_eqn = apfst (snd o strip_comb) o Logic.dest_equals o Logic.unvarify_global;
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   965
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
   966
abstype cert = Nothing of thm
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
   967
  | Equations of thm * bool list
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   968
  | Projection of term * string
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   969
  | Abstract of thm * string
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
   970
with
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34874
diff changeset
   971
55364
4d26690379b1 build up preprocessing context only once
haftmann
parents: 55363
diff changeset
   972
fun dummy_thm ctxt c =
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34874
diff changeset
   973
  let
55364
4d26690379b1 build up preprocessing context only once
haftmann
parents: 55363
diff changeset
   974
    val thy = Proof_Context.theory_of ctxt;
54603
89d5b69e5a08 prefer name-normalizing devarify over unvarifyT whenever appropriate
haftmann
parents: 53171
diff changeset
   975
    val raw_ty = devarify (const_typ thy c);
49534
791e6fc53f73 more strict typscheme_equiv check: must fix variables of more specific type;
haftmann
parents: 49533
diff changeset
   976
    val (vs, _) = typscheme thy (c, raw_ty);
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 51584
diff changeset
   977
    val sortargs = case Axclass.class_of_param thy c
40761
1ef64dcb24b7 corrected: use canonical variables of type scheme uniformly
haftmann
parents: 40758
diff changeset
   978
     of SOME class => [[class]]
1ef64dcb24b7 corrected: use canonical variables of type scheme uniformly
haftmann
parents: 40758
diff changeset
   979
      | NONE => (case get_type_of_constr_or_abstr thy c
1ef64dcb24b7 corrected: use canonical variables of type scheme uniformly
haftmann
parents: 40758
diff changeset
   980
         of SOME (tyco, _) => (map snd o fst o the)
1ef64dcb24b7 corrected: use canonical variables of type scheme uniformly
haftmann
parents: 40758
diff changeset
   981
              (AList.lookup (op =) ((snd o fst o get_type thy) tyco) c)
1ef64dcb24b7 corrected: use canonical variables of type scheme uniformly
haftmann
parents: 40758
diff changeset
   982
          | NONE => replicate (length vs) []);
1ef64dcb24b7 corrected: use canonical variables of type scheme uniformly
haftmann
parents: 40758
diff changeset
   983
    val the_sort = the o AList.lookup (op =) (map fst vs ~~ sortargs);
1ef64dcb24b7 corrected: use canonical variables of type scheme uniformly
haftmann
parents: 40758
diff changeset
   984
    val ty = map_type_tfree (fn (v, _) => TFree (v, the_sort v)) raw_ty
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   985
    val chead = build_head thy (c, ty);
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
   986
  in Thm.weaken chead Drule.dummy_thm end;
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34874
diff changeset
   987
55364
4d26690379b1 build up preprocessing context only once
haftmann
parents: 55363
diff changeset
   988
fun nothing_cert ctxt c = Nothing (dummy_thm ctxt c);
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
   989
55364
4d26690379b1 build up preprocessing context only once
haftmann
parents: 55363
diff changeset
   990
fun cert_of_eqns ctxt c [] = Equations (dummy_thm ctxt c, [])
66167
1bd268ab885c more informative task_statistics;
wenzelm
parents: 66149
diff changeset
   991
  | cert_of_eqns ctxt c raw_eqns =
34874
89c230bf8feb added code certificates
haftmann
parents: 34272
diff changeset
   992
      let
55364
4d26690379b1 build up preprocessing context only once
haftmann
parents: 55363
diff changeset
   993
        val thy = Proof_Context.theory_of ctxt;
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   994
        val eqns = burrow_fst (canonize_thms thy) raw_eqns;
63240
f82c0b803bda more correct exception handling
haftmann
parents: 63239
diff changeset
   995
        val _ = map (assert_eqn thy) eqns;
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34874
diff changeset
   996
        val (thms, propers) = split_list eqns;
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   997
        val _ = map (fn thm => if c = const_eqn thy thm then ()
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
   998
          else error ("Wrong head of code equation,\nexpected constant "
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 60949
diff changeset
   999
            ^ string_of_const thy c ^ "\n" ^ Thm.string_of_thm_global thy thm)) thms;
74235
dbaed92fd8af tuned signature;
wenzelm
parents: 73968
diff changeset
  1000
        val tvars_of = build_rev o Term.add_tvarsT;
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34874
diff changeset
  1001
        val vss = map (tvars_of o snd o head_eqn) thms;
75353
05f7f5454cb6 prefer build combinator
haftmann
parents: 74978
diff changeset
  1002
        val inter_sorts =
05f7f5454cb6 prefer build combinator
haftmann
parents: 74978
diff changeset
  1003
          build o fold (curry (Sorts.inter_sort (Sign.classes_of thy)) o snd);
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34874
diff changeset
  1004
        val sorts = map_transpose inter_sorts vss;
81507
08574da77b4a clarified signature: shorten common cases;
wenzelm
parents: 79411
diff changeset
  1005
        val vts = Name.invent_types_global sorts;
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74235
diff changeset
  1006
        fun instantiate vs =
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74235
diff changeset
  1007
          Thm.instantiate (TVars.make (vs ~~ map (Thm.ctyp_of ctxt o TFree) vts), Vars.empty);
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74235
diff changeset
  1008
        val thms' = map2 instantiate vss thms;
40758
4f2c3e842ef8 consider sort constraints for datatype constructors when constructing the empty equation certificate;
haftmann
parents: 40726
diff changeset
  1009
        val head_thm = Thm.symmetric (Thm.assume (build_head thy (head_eqn (hd thms'))));
34874
89c230bf8feb added code certificates
haftmann
parents: 34272
diff changeset
  1010
        fun head_conv ct = if can Thm.dest_comb ct
89c230bf8feb added code certificates
haftmann
parents: 34272
diff changeset
  1011
          then Conv.fun_conv head_conv ct
89c230bf8feb added code certificates
haftmann
parents: 34272
diff changeset
  1012
          else Conv.rewr_conv head_thm ct;
89c230bf8feb added code certificates
haftmann
parents: 34272
diff changeset
  1013
        val rewrite_head = Conv.fconv_rule (Conv.arg1_conv head_conv);
40758
4f2c3e842ef8 consider sort constraints for datatype constructors when constructing the empty equation certificate;
haftmann
parents: 40726
diff changeset
  1014
        val cert_thm = Conjunction.intr_balanced (map rewrite_head thms');
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1015
      in Equations (cert_thm, propers) end;
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34874
diff changeset
  1016
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1017
fun cert_of_proj ctxt proj tyco =
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1018
  let
66127
0561ac527270 tuned internal signature
haftmann
parents: 66126
diff changeset
  1019
    val thy = Proof_Context.theory_of ctxt
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1020
    val (vs, {abstractor = (abs, (_, ty)), projection = proj', ...}) = get_abstype_spec thy tyco;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1021
    val _ = if proj = proj' then () else
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1022
      error ("Wrong head of projection,\nexpected constant " ^ string_of_const thy proj);
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1023
  in Projection (mk_proj tyco vs ty abs proj, tyco) end;
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1024
66127
0561ac527270 tuned internal signature
haftmann
parents: 66126
diff changeset
  1025
fun cert_of_abs ctxt tyco c raw_abs_thm =
34874
89c230bf8feb added code certificates
haftmann
parents: 34272
diff changeset
  1026
  let
66127
0561ac527270 tuned internal signature
haftmann
parents: 66126
diff changeset
  1027
    val thy = Proof_Context.theory_of ctxt;
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1028
    val abs_thm = singleton (canonize_thms thy) raw_abs_thm;
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1029
    val _ = assert_abs_eqn thy (SOME tyco) abs_thm;
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1030
    val _ = if c = const_abs_eqn thy abs_thm then ()
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1031
      else error ("Wrong head of abstract code equation,\nexpected constant "
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 60949
diff changeset
  1032
        ^ string_of_const thy c ^ "\n" ^ Thm.string_of_thm_global thy abs_thm);
36615
88756a5a92fc renamed Thm.freezeT to Thm.legacy_freezeT -- it is based on Type.legacy_freeze;
wenzelm
parents: 36610
diff changeset
  1033
  in Abstract (Thm.legacy_freezeT abs_thm, tyco) end;
34874
89c230bf8feb added code certificates
haftmann
parents: 34272
diff changeset
  1034
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1035
fun constrain_cert_thm thy sorts cert_thm =
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1036
  let
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1037
    val ((vs, _), head) = get_head thy cert_thm;
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1038
    val (subst, cert_thm') = cert_thm
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1039
      |> Thm.implies_intr head
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1040
      |> constrain_thm thy vs sorts;
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1041
    val head' = Thm.term_of head
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1042
      |> subst
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59458
diff changeset
  1043
      |> Thm.global_cterm_of thy;
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1044
    val cert_thm'' = cert_thm'
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1045
      |> Thm.elim_implies (Thm.assume head');
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1046
  in cert_thm'' end;
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1047
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1048
fun constrain_cert thy sorts (Nothing cert_thm) =
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1049
      Nothing (constrain_cert_thm thy sorts cert_thm)
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1050
  | constrain_cert thy sorts (Equations (cert_thm, propers)) =
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1051
      Equations (constrain_cert_thm thy sorts cert_thm, propers)
66324
haftmann
parents: 66310
diff changeset
  1052
  | constrain_cert _ _ (cert as Projection _) =
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1053
      cert
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1054
  | constrain_cert thy sorts (Abstract (abs_thm, tyco)) =
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1055
      Abstract (snd (constrain_thm thy (fst (typscheme_abs thy abs_thm)) sorts abs_thm), tyco);
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1056
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1057
fun conclude_cert (Nothing cert_thm) =
78056
904242f46ce1 more accurate Thm.trim_context / Thm.transfer;
wenzelm
parents: 78054
diff changeset
  1058
      Nothing (Thm.close_derivation \<^here> cert_thm |> Thm.trim_context)
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1059
  | conclude_cert (Equations (cert_thm, propers)) =
78056
904242f46ce1 more accurate Thm.trim_context / Thm.transfer;
wenzelm
parents: 78054
diff changeset
  1060
      Equations (Thm.close_derivation \<^here> cert_thm |> Thm.trim_context, propers)
49971
8b50286c36d3 close code theorems explicitly after preprocessing
haftmann
parents: 49904
diff changeset
  1061
  | conclude_cert (cert as Projection _) =
8b50286c36d3 close code theorems explicitly after preprocessing
haftmann
parents: 49904
diff changeset
  1062
      cert
8b50286c36d3 close code theorems explicitly after preprocessing
haftmann
parents: 49904
diff changeset
  1063
  | conclude_cert (Abstract (abs_thm, tyco)) =
78056
904242f46ce1 more accurate Thm.trim_context / Thm.transfer;
wenzelm
parents: 78054
diff changeset
  1064
      Abstract (Thm.close_derivation \<^here> abs_thm |> Thm.trim_context, tyco);
49971
8b50286c36d3 close code theorems explicitly after preprocessing
haftmann
parents: 49904
diff changeset
  1065
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1066
fun typscheme_of_cert thy (Nothing cert_thm) =
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1067
      fst (get_head thy cert_thm)
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1068
  | typscheme_of_cert thy (Equations (cert_thm, _)) =
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1069
      fst (get_head thy cert_thm)
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1070
  | typscheme_of_cert thy (Projection (proj, _)) =
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1071
      typscheme_projection thy proj
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1072
  | typscheme_of_cert thy (Abstract (abs_thm, _)) =
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1073
      typscheme_abs thy abs_thm;
34874
89c230bf8feb added code certificates
haftmann
parents: 34272
diff changeset
  1074
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1075
fun typargs_deps_of_cert thy (Nothing cert_thm) =
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1076
      let
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1077
        val vs = (fst o fst) (get_head thy cert_thm);
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1078
      in (vs, []) end
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1079
  | typargs_deps_of_cert thy (Equations (cert_thm, propers)) =
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1080
      let
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1081
        val vs = (fst o fst) (get_head thy cert_thm);
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1082
        val equations = if null propers then [] else
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1083
          Thm.prop_of cert_thm
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1084
          |> Logic.dest_conjunction_balanced (length propers);
75353
05f7f5454cb6 prefer build combinator
haftmann
parents: 74978
diff changeset
  1085
      in (vs, build (fold (add_rhss_of_eqn thy) equations)) end
40758
4f2c3e842ef8 consider sort constraints for datatype constructors when constructing the empty equation certificate;
haftmann
parents: 40726
diff changeset
  1086
  | typargs_deps_of_cert thy (Projection (t, _)) =
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1087
      (fst (typscheme_projection thy t), add_rhss_of_eqn thy t [])
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1088
  | typargs_deps_of_cert thy (Abstract (abs_thm, tyco)) =
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1089
      let
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1090
        val vs = fst (typscheme_abs thy abs_thm);
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1091
        val (_, concrete_thm) = concretify_abs thy tyco abs_thm;
45344
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 45211
diff changeset
  1092
      in (vs, add_rhss_of_eqn thy (Logic.unvarify_types_global (Thm.prop_of concrete_thm)) []) end;
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
  1093
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1094
fun equations_of_cert thy (cert as Nothing _) =
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1095
      (typscheme_of_cert thy cert, NONE)
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1096
  | equations_of_cert thy (cert as Equations (cert_thm, propers)) =
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1097
      let
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1098
        val tyscm = typscheme_of_cert thy cert;
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1099
        val thms = if null propers then [] else
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1100
          cert_thm
78056
904242f46ce1 more accurate Thm.trim_context / Thm.transfer;
wenzelm
parents: 78054
diff changeset
  1101
          |> Thm.transfer thy
35624
c4e29a0bb8c1 modernized structure Local_Defs;
wenzelm
parents: 35376
diff changeset
  1102
          |> Local_Defs.expand [snd (get_head thy cert_thm)]
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35624
diff changeset
  1103
          |> Thm.varifyT_global
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1104
          |> Conjunction.elim_balanced (length propers);
77702
haftmann
parents: 75399
diff changeset
  1105
        fun abstractions (args, rhs) = (map (pair NONE) args, (NONE, rhs));
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1106
      in (tyscm, SOME (map (abstractions o dest_eqn o Thm.prop_of) thms ~~ (map SOME thms ~~ propers))) end
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1107
  | equations_of_cert thy (Projection (t, tyco)) =
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1108
      let
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1109
        val (_, {abstractor = (abs, _), ...}) = get_abstype_spec thy tyco;
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1110
        val tyscm = typscheme_projection thy t;
45344
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 45211
diff changeset
  1111
        val t' = Logic.varify_types_global t;
77702
haftmann
parents: 75399
diff changeset
  1112
        fun abstractions (args, rhs) = (map (pair (SOME abs)) args, (NONE, rhs));
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1113
      in (tyscm, SOME [((abstractions o dest_eqn) t', (NONE, true))]) end
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1114
  | equations_of_cert thy (Abstract (abs_thm, tyco)) =
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1115
      let
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1116
        val tyscm = typscheme_abs thy abs_thm;
78056
904242f46ce1 more accurate Thm.trim_context / Thm.transfer;
wenzelm
parents: 78054
diff changeset
  1117
        val (abs, concrete_thm) = concretify_abs thy tyco (Thm.transfer thy abs_thm);
77702
haftmann
parents: 75399
diff changeset
  1118
        fun abstractions (args, rhs) = (map (pair NONE) args, (SOME abs, rhs));
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35624
diff changeset
  1119
      in
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1120
        (tyscm, SOME [((abstractions o dest_eqn o Thm.prop_of) concrete_thm,
36209
566be5d48090 more appropriate representation of valid positions for abstractors; more accurate checking of abstype certificates
haftmann
parents: 36202
diff changeset
  1121
          (SOME (Thm.varifyT_global abs_thm), true))])
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35624
diff changeset
  1122
      end;
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
  1123
66324
haftmann
parents: 66310
diff changeset
  1124
fun pretty_cert _ (Nothing _) =
66332
489667636064 compactified output
haftmann
parents: 66330
diff changeset
  1125
      []
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1126
  | pretty_cert thy (cert as Equations _) =
63073
413184c7a2a2 clarified context, notably for internal use of Simplifier;
wenzelm
parents: 61268
diff changeset
  1127
      (map_filter
413184c7a2a2 clarified context, notably for internal use of Simplifier;
wenzelm
parents: 61268
diff changeset
  1128
        (Option.map (Thm.pretty_thm_global thy o
413184c7a2a2 clarified context, notably for internal use of Simplifier;
wenzelm
parents: 61268
diff changeset
  1129
            Axclass.overload (Proof_Context.init_global thy)) o fst o snd)
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1130
         o these o snd o equations_of_cert thy) cert
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1131
  | pretty_cert thy (Projection (t, _)) =
45344
e209da839ff4 added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents: 45211
diff changeset
  1132
      [Syntax.pretty_term_global thy (Logic.varify_types_global t)]
40758
4f2c3e842ef8 consider sort constraints for datatype constructors when constructing the empty equation certificate;
haftmann
parents: 40726
diff changeset
  1133
  | pretty_cert thy (Abstract (abs_thm, _)) =
63073
413184c7a2a2 clarified context, notably for internal use of Simplifier;
wenzelm
parents: 61268
diff changeset
  1134
      [(Thm.pretty_thm_global thy o
413184c7a2a2 clarified context, notably for internal use of Simplifier;
wenzelm
parents: 61268
diff changeset
  1135
         Axclass.overload (Proof_Context.init_global thy) o Thm.varifyT_global) abs_thm];
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1136
34895
19fd499cddff explicit abstract type of code certificates
haftmann
parents: 34894
diff changeset
  1137
end;
34891
99b9a6290446 code certificates as integral part of code generation
haftmann
parents: 34874
diff changeset
  1138
34874
89c230bf8feb added code certificates
haftmann
parents: 34272
diff changeset
  1139
57430
020cea57eaa4 tracing facilities for the code generator preprocessor
haftmann
parents: 57429
diff changeset
  1140
(* code certificate access with preprocessing *)
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1141
48075
ec5e62b868eb apply preprocessing simpset also to rhs of abstract code equations
haftmann
parents: 48068
diff changeset
  1142
fun eqn_conv conv ct =
ec5e62b868eb apply preprocessing simpset also to rhs of abstract code equations
haftmann
parents: 48068
diff changeset
  1143
  let
ec5e62b868eb apply preprocessing simpset also to rhs of abstract code equations
haftmann
parents: 48068
diff changeset
  1144
    fun lhs_conv ct = if can Thm.dest_comb ct
ec5e62b868eb apply preprocessing simpset also to rhs of abstract code equations
haftmann
parents: 48068
diff changeset
  1145
      then Conv.combination_conv lhs_conv conv ct
ec5e62b868eb apply preprocessing simpset also to rhs of abstract code equations
haftmann
parents: 48068
diff changeset
  1146
      else Conv.all_conv ct;
ec5e62b868eb apply preprocessing simpset also to rhs of abstract code equations
haftmann
parents: 48068
diff changeset
  1147
  in Conv.combination_conv (Conv.arg_conv lhs_conv) conv ct end;
ec5e62b868eb apply preprocessing simpset also to rhs of abstract code equations
haftmann
parents: 48068
diff changeset
  1148
55364
4d26690379b1 build up preprocessing context only once
haftmann
parents: 55363
diff changeset
  1149
fun rewrite_eqn conv ctxt =
4d26690379b1 build up preprocessing context only once
haftmann
parents: 55363
diff changeset
  1150
  singleton (Variable.trade (K (map (Conv.fconv_rule (conv (Simplifier.rewrite ctxt))))) ctxt)
48075
ec5e62b868eb apply preprocessing simpset also to rhs of abstract code equations
haftmann
parents: 48068
diff changeset
  1151
74978
1f30aa91f496 more correct transfer
haftmann
parents: 74561
diff changeset
  1152
fun apply_functrans ctxt functrans =
58559
d230e7075bcf code preprocessor tracing also for function transformers
haftmann
parents: 58011
diff changeset
  1153
  let
82588
wenzelm
parents: 82587
diff changeset
  1154
    fun trace_eqns s eqns =
wenzelm
parents: 82587
diff changeset
  1155
      Pretty.writeln (Pretty.chunks
wenzelm
parents: 82587
diff changeset
  1156
        (Pretty.str s :: map (Thm.pretty_thm ctxt o fst) eqns));
58559
d230e7075bcf code preprocessor tracing also for function transformers
haftmann
parents: 58011
diff changeset
  1157
    val tracing = if Config.get ctxt simp_trace then trace_eqns else (K o K) ();
d230e7075bcf code preprocessor tracing also for function transformers
haftmann
parents: 58011
diff changeset
  1158
  in
d230e7075bcf code preprocessor tracing also for function transformers
haftmann
parents: 58011
diff changeset
  1159
    tap (tracing "before function transformation")
d230e7075bcf code preprocessor tracing also for function transformers
haftmann
parents: 58011
diff changeset
  1160
    #> (perhaps o perhaps_loop o perhaps_apply) functrans
d230e7075bcf code preprocessor tracing also for function transformers
haftmann
parents: 58011
diff changeset
  1161
    #> tap (tracing "after function transformation")
d230e7075bcf code preprocessor tracing also for function transformers
haftmann
parents: 58011
diff changeset
  1162
  end;
48075
ec5e62b868eb apply preprocessing simpset also to rhs of abstract code equations
haftmann
parents: 48068
diff changeset
  1163
74978
1f30aa91f496 more correct transfer
haftmann
parents: 74561
diff changeset
  1164
fun preprocess conv ctxt =
1f30aa91f496 more correct transfer
haftmann
parents: 74561
diff changeset
  1165
  rewrite_eqn conv ctxt
1f30aa91f496 more correct transfer
haftmann
parents: 74561
diff changeset
  1166
  #> Axclass.unoverload ctxt;
1f30aa91f496 more correct transfer
haftmann
parents: 74561
diff changeset
  1167
57429
4aef934d43ad tuned interface
haftmann
parents: 56811
diff changeset
  1168
fun get_cert ctxt functrans c =
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1169
  case lookup_proper_fun_spec (specs_of (Proof_Context.theory_of ctxt)) c of
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1170
    NONE => nothing_cert ctxt c
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1171
  | SOME (Eqns (_, eqns)) => eqns
74978
1f30aa91f496 more correct transfer
haftmann
parents: 74561
diff changeset
  1172
      |> (map o apfst) (Thm.transfer' ctxt)
1f30aa91f496 more correct transfer
haftmann
parents: 74561
diff changeset
  1173
      |> apply_functrans ctxt functrans
1f30aa91f496 more correct transfer
haftmann
parents: 74561
diff changeset
  1174
      |> (map o apfst) (preprocess eqn_conv ctxt)
1f30aa91f496 more correct transfer
haftmann
parents: 74561
diff changeset
  1175
      |> cert_of_eqns ctxt c
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1176
  | SOME (Proj (_, (tyco, _))) => cert_of_proj ctxt c tyco
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1177
  | SOME (Abstr (abs_thm, (tyco, _))) => abs_thm
74978
1f30aa91f496 more correct transfer
haftmann
parents: 74561
diff changeset
  1178
      |> Thm.transfer' ctxt
1f30aa91f496 more correct transfer
haftmann
parents: 74561
diff changeset
  1179
      |> preprocess Conv.arg_conv ctxt
1f30aa91f496 more correct transfer
haftmann
parents: 74561
diff changeset
  1180
      |> cert_of_abs ctxt tyco c;
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
  1181
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
  1182
66128
0181bb24bdca more uniform ordering and naming of sections;
haftmann
parents: 66127
diff changeset
  1183
(* case certificates *)
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1184
66128
0181bb24bdca more uniform ordering and naming of sections;
haftmann
parents: 66127
diff changeset
  1185
local
0181bb24bdca more uniform ordering and naming of sections;
haftmann
parents: 66127
diff changeset
  1186
0181bb24bdca more uniform ordering and naming of sections;
haftmann
parents: 66127
diff changeset
  1187
fun raw_case_cert thm =
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1188
  let
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1189
    val ((head, raw_case_expr), cases) = (apfst Logic.dest_equals
32640
ba6531df2c64 corrected order of type variables in code equations; more precise certificate for cases
haftmann
parents: 32544
diff changeset
  1190
      o apsnd Logic.dest_conjunctions o Logic.dest_implies o Thm.plain_prop_of) thm;
66132
5844a096c462 dropped void values
haftmann
parents: 66131
diff changeset
  1191
    val _ = case head of Free _ => ()
5844a096c462 dropped void values
haftmann
parents: 66131
diff changeset
  1192
      | Var _ => ()
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1193
      | _ => raise TERM ("case_cert", []);
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1194
    val ([(case_var, _)], case_expr) = Term.strip_abs_eta 1 raw_case_expr;
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1195
    val (Const (case_const, _), raw_params) = strip_comb case_expr;
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1196
    val n = find_index (fn Free (v, _) => v = case_var | _ => false) raw_params;
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1197
    val _ = if n = ~1 then raise TERM ("case_cert", []) else ();
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1198
    val params = map (fst o dest_Var) (nth_drop n raw_params);
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1199
    fun dest_case t =
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1200
      let
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1201
        val (head' $ t_co, rhs) = Logic.dest_equals t;
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1202
        val _ = if head' = head then () else raise TERM ("case_cert", []);
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1203
        val (Const (co, _), args) = strip_comb t_co;
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1204
        val (Var (param, _), args') = strip_comb rhs;
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1205
        val _ = if args' = args then () else raise TERM ("case_cert", []);
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1206
      in (param, co) end;
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1207
    fun analyze_cases cases =
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1208
      let
75353
05f7f5454cb6 prefer build combinator
haftmann
parents: 74978
diff changeset
  1209
        val co_list = build (fold (AList.update (op =) o dest_case) cases);
47437
4625ee486ff6 generalise case certificates to allow ignored parameters
Andreas Lochbihler
parents: 46513
diff changeset
  1210
      in map (AList.lookup (op =) co_list) params end;
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1211
    fun analyze_let t =
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1212
      let
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1213
        val (head' $ arg, Var (param', _) $ arg') = Logic.dest_equals t;
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1214
        val _ = if head' = head then () else raise TERM ("case_cert", []);
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1215
        val _ = if arg' = arg then () else raise TERM ("case_cert", []);
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1216
        val _ = if [param'] = params then () else raise TERM ("case_cert", []);
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1217
      in [] end;
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1218
    fun analyze (cases as [let_case]) =
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1219
          (analyze_cases cases handle Bind => analyze_let let_case)
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1220
      | analyze cases = analyze_cases cases;
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1221
  in (case_const, (n, analyze cases)) end;
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1222
66128
0181bb24bdca more uniform ordering and naming of sections;
haftmann
parents: 66127
diff changeset
  1223
in
0181bb24bdca more uniform ordering and naming of sections;
haftmann
parents: 66127
diff changeset
  1224
0181bb24bdca more uniform ordering and naming of sections;
haftmann
parents: 66127
diff changeset
  1225
fun case_cert thm = raw_case_cert thm
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1226
  handle Bind => error "bad case certificate"
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1227
       | TERM _ => error "bad case certificate";
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31152
diff changeset
  1228
66128
0181bb24bdca more uniform ordering and naming of sections;
haftmann
parents: 66127
diff changeset
  1229
end;
0181bb24bdca more uniform ordering and naming of sections;
haftmann
parents: 66127
diff changeset
  1230
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1231
fun lookup_case_spec thy = History.lookup ((cases_of o specs_of) thy);
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1232
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1233
fun get_case_schema thy c = case lookup_case_spec thy c of
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1234
    SOME (Case {schema, ...}) => SOME schema
66189
23917e861eaa treat "undefined" constants internally as special form of case combinators
haftmann
parents: 66167
diff changeset
  1235
  | _ => NONE;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1236
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1237
fun get_case_cong thy c = case lookup_case_spec thy c of
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1238
    SOME (Case {cong, ...}) => SOME cong
66189
23917e861eaa treat "undefined" constants internally as special form of case combinators
haftmann
parents: 66167
diff changeset
  1239
  | _ => NONE;
23917e861eaa treat "undefined" constants internally as special form of case combinators
haftmann
parents: 66167
diff changeset
  1240
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1241
fun is_undefined thy c = case lookup_case_spec thy c of
66189
23917e861eaa treat "undefined" constants internally as special form of case combinators
haftmann
parents: 66167
diff changeset
  1242
    SOME Undefined => true
23917e861eaa treat "undefined" constants internally as special form of case combinators
haftmann
parents: 66167
diff changeset
  1243
  | _ => false;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1244
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1245
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
  1246
(* diagnostic *)
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1247
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1248
fun print_codesetup thy =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1249
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
  1250
    val ctxt = Proof_Context.init_global thy;
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1251
    val specs = specs_of thy;
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1252
    fun pretty_equations const thms =
51584
98029ceda8ce more item markup;
wenzelm
parents: 51580
diff changeset
  1253
      (Pretty.block o Pretty.fbreaks)
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 60949
diff changeset
  1254
        (Pretty.str (string_of_const thy const) :: map (Thm.pretty_thm_item ctxt) thms);
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1255
    fun pretty_function (const, Eqns (_, eqns)) =
66131
39e1c876bfec more uniform order of constructors
haftmann
parents: 66130
diff changeset
  1256
          pretty_equations const (map fst eqns)
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1257
      | pretty_function (const, Proj (proj, _)) = Pretty.block
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1258
          [Pretty.str (string_of_const thy const), Pretty.fbrk, Syntax.pretty_term ctxt proj]
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1259
      | pretty_function (const, Abstr (thm, _)) = pretty_equations const [thm];
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1260
    fun pretty_typ (tyco, vs) = Pretty.str
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1261
      (string_of_typ thy (Type (tyco, map TFree vs)));
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1262
    fun pretty_type_spec (typ, (cos, abstract)) = if null cos
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1263
      then pretty_typ typ
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1264
      else (Pretty.block o Pretty.breaks) (
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1265
        pretty_typ typ
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1266
        :: Pretty.str "="
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1267
        :: (if abstract then [Pretty.str "(abstract)"] else [])
40726
16dcfedc4eb7 keep type variable arguments of datatype constructors in bookkeeping
haftmann
parents: 40564
diff changeset
  1268
        @ separate (Pretty.str "|") (map (fn (c, (_, [])) => Pretty.str (string_of_const thy c)
16dcfedc4eb7 keep type variable arguments of datatype constructors in bookkeeping
haftmann
parents: 40564
diff changeset
  1269
             | (c, (_, tys)) =>
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1270
                 (Pretty.block o Pretty.breaks)
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1271
                    (Pretty.str (string_of_const thy c)
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1272
                      :: Pretty.str "of"
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1273
                      :: map (Pretty.quote o Syntax.pretty_typ_global thy) tys)) cos)
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1274
      );
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1275
    fun pretty_case_param NONE = "<ignored>"
75399
cdf84288d93c pass constructor arity as part of case certficiate
haftmann
parents: 75353
diff changeset
  1276
      | pretty_case_param (SOME (c, _)) = string_of_const thy c
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1277
    fun pretty_case (const, Case {schema = (_, (_, [])), ...}) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1278
          Pretty.str (string_of_const thy const)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1279
      | pretty_case (const, Case {schema = (_, (_, cos)), ...}) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1280
          (Pretty.block o Pretty.breaks) [
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1281
            Pretty.str (string_of_const thy const), Pretty.str "with",
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1282
            (Pretty.block o Pretty.commas o map (Pretty.str o pretty_case_param)) cos]
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1283
      | pretty_case (const, Undefined) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1284
          (Pretty.block o Pretty.breaks) [
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1285
            Pretty.str (string_of_const thy const), Pretty.str "<undefined>"];
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1286
    val functions = all_fun_specs specs
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58666
diff changeset
  1287
      |> sort (string_ord o apply2 fst);
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1288
    val types = History.all (types_of specs)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1289
      |> map (fn (tyco, (vs, spec)) =>
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1290
          ((tyco, vs), constructors_of spec))
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58666
diff changeset
  1291
      |> sort (string_ord o apply2 (fst o fst));
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1292
    val cases = History.all (cases_of specs)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1293
      |> filter (fn (_, No_Case) => false | _ => true)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1294
      |> sort (string_ord o apply2 fst);
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1295
  in
82587
7415414bd9d8 more scalable: discontinue odd shortcuts from 6b3739fee456, which produce bulky strings internally;
wenzelm
parents: 81521
diff changeset
  1296
    Pretty.writeln (Pretty.chunks [
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1297
      Pretty.block (
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1298
        Pretty.str "types:" :: Pretty.fbrk
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1299
        :: (Pretty.fbreaks o map pretty_type_spec) types
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1300
      ),
25968
66cfe1d00be0 print postprocessor equations
haftmann
parents: 25597
diff changeset
  1301
      Pretty.block (
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1302
        Pretty.str "functions:" :: Pretty.fbrk
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1303
        :: (Pretty.fbreaks o map pretty_function) functions
34901
0d6a2ae86525 printing of cases
haftmann
parents: 34895
diff changeset
  1304
      ),
0d6a2ae86525 printing of cases
haftmann
parents: 34895
diff changeset
  1305
      Pretty.block (
0d6a2ae86525 printing of cases
haftmann
parents: 34895
diff changeset
  1306
        Pretty.str "cases:" :: Pretty.fbrk
0d6a2ae86525 printing of cases
haftmann
parents: 34895
diff changeset
  1307
        :: (Pretty.fbreaks o map pretty_case) cases
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1308
      )
82587
7415414bd9d8 more scalable: discontinue odd shortcuts from 6b3739fee456, which produce bulky strings internally;
wenzelm
parents: 81521
diff changeset
  1309
    ])
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1310
  end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1311
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1312
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1313
(** declaration of executable ingredients **)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1314
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1315
(* plugins for dependent applications *)
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1316
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1317
structure Codetype_Plugin = Plugin(type T = string);
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1318
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 67032
diff changeset
  1319
val codetype_plugin = Plugin_Name.declare_setup \<^binding>\<open>codetype\<close>;
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1320
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1321
fun type_interpretation f =
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1322
  Codetype_Plugin.interpretation codetype_plugin
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1323
    (fn tyco => Local_Theory.background_theory
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1324
      (fn thy =>
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1325
        thy
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1326
        |> Sign.root_path
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1327
        |> Sign.add_path (Long_Name.qualifier tyco)
67032
ed499d1252fc strip some trailing spaces to force Pure rebuild after ce6454669360
Lars Hupel <lars.hupel@mytum.de>
parents: 66332
diff changeset
  1328
        |> f tyco
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1329
        |> Sign.restore_naming thy));
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1331
fun datatype_interpretation f =
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1332
  type_interpretation (fn tyco => fn thy =>
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1333
    case get_type thy tyco of
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1334
      (spec, false) => f (tyco, spec) thy
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1335
    | (_, true) => thy
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1336
  );
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1337
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1338
fun abstype_interpretation f =
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1339
  type_interpretation (fn tyco => fn thy =>
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1340
    case try (get_abstype_spec thy) tyco of
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1341
      SOME spec => f (tyco, spec) thy
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1342
    | NONE => thy
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1343
  );
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1344
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1345
fun register_tyco_for_plugin tyco =
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1346
  Named_Target.theory_map (Codetype_Plugin.data_default tyco);
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1347
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1348
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1349
(* abstract code declarations *)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1350
78140
90a43a9b3605 proper trim_context;
wenzelm
parents: 78139
diff changeset
  1351
fun code_declaration (strictness: strictness) transform f x =
90a43a9b3605 proper trim_context;
wenzelm
parents: 78139
diff changeset
  1352
  let val x0 = transform Morphism.trim_context_morphism x in
90a43a9b3605 proper trim_context;
wenzelm
parents: 78139
diff changeset
  1353
    Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>}
90a43a9b3605 proper trim_context;
wenzelm
parents: 78139
diff changeset
  1354
      (fn phi => Context.mapping (f strictness (transform phi x0)) I)
90a43a9b3605 proper trim_context;
wenzelm
parents: 78139
diff changeset
  1355
  end;
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1356
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1357
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1358
(* types *)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1359
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1360
fun invalidate_constructors_of (_, type_spec) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1361
  fold (fn (c, _) => History.register c unimplemented) (fst (constructors_of type_spec));
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1362
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1363
fun invalidate_abstract_functions_of (_, type_spec) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1364
  fold (fn c => History.register c unimplemented) (abstract_functions_of type_spec);
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1365
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1366
fun invalidate_case_combinators_of (_, type_spec) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1367
  fold (fn c => History.register c No_Case) (case_combinators_of type_spec);
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1368
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1369
fun register_type (tyco, vs_typ_spec) specs =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1370
  let
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1371
    val olds = the_list (History.lookup (types_of specs) tyco);
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1372
  in
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1373
    specs
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1374
    |> map_functions (fold invalidate_abstract_functions_of olds
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1375
        #> invalidate_constructors_of vs_typ_spec)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1376
    |> map_cases (fold invalidate_case_combinators_of olds)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1377
    |> map_types (History.register tyco vs_typ_spec)
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1378
  end;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1379
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1380
fun declare_datatype_global proto_constrs thy =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1381
  let
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1382
    fun unoverload_const_typ (c, ty) =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1383
      (Axclass.unoverload_const thy (c, ty), ty);
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1384
    val constrs = map unoverload_const_typ proto_constrs;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1385
    val (tyco, (vs, cos)) = constrset_of_consts thy constrs;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1386
  in
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1387
    thy
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1388
    |> modify_specs (register_type
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1389
        (tyco, (vs, Constructors {constructors = cos, case_combinators = []})))
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1390
    |> register_tyco_for_plugin tyco
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1391
  end;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1392
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1393
fun declare_datatype_cmd raw_constrs thy =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1394
  declare_datatype_global (map (read_bare_const thy) raw_constrs) thy;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1395
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1396
fun generic_declare_abstype strictness proto_thm thy =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1397
  case check_abstype_cert strictness thy proto_thm of
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1398
    SOME (tyco, (vs, (abstractor as (abs, (_, ty)), (proj, abs_rep)))) =>
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1399
      thy
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1400
      |> modify_specs (register_type
78056
904242f46ce1 more accurate Thm.trim_context / Thm.transfer;
wenzelm
parents: 78054
diff changeset
  1401
            (tyco, (vs, Abstractor {abstractor = abstractor, projection = proj,
904242f46ce1 more accurate Thm.trim_context / Thm.transfer;
wenzelm
parents: 78054
diff changeset
  1402
                abs_rep = Thm.trim_context abs_rep, more_abstract_functions = []}))
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1403
          #> register_fun_spec proj
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1404
            (Proj (Logic.varify_types_global (mk_proj tyco vs ty abs proj), (tyco, abs))))
66330
dcb3e6bdc00a one single plugin for code type declarations avoids problems when bootstrapping new plugins over types which have been both declared concrete and abstract in their code historiy
haftmann
parents: 66329
diff changeset
  1405
      |> register_tyco_for_plugin tyco
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1406
  | NONE => thy;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1407
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1408
val declare_abstype_global = generic_declare_abstype Strict;
78068
a01c3bcf22dd tuned: avoid duplication;
wenzelm
parents: 78056
diff changeset
  1409
val declare_abstype = code_declaration Liberal Morphism.thm generic_declare_abstype;
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1410
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1411
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1412
(* functions *)
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
  1413
63239
d562c9948dee explicit tagging of code equations de-baroquifies interface
haftmann
parents: 63238
diff changeset
  1414
(*
63242
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
  1415
  strictness wrt. shape of theorem propositions:
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1416
    * default equations: silent
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1417
    * using declarations and attributes: warnings (after morphism application!)
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1418
    * using global declarations (... -> thy -> thy): strict
63242
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
  1419
    * internal processing after storage: strict
63239
d562c9948dee explicit tagging of code equations de-baroquifies interface
haftmann
parents: 63238
diff changeset
  1420
*)
d562c9948dee explicit tagging of code equations de-baroquifies interface
haftmann
parents: 63238
diff changeset
  1421
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1422
local
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1423
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1424
fun subsumptive_append thy { verbose } (thm, proper) eqns =
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1425
  let
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1426
    val args_of = drop_prefix is_Var o rev o snd o strip_comb
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1427
      o Term.strip_sorts o fst o Logic.dest_equals o Thm.plain_prop_of
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1428
      o Thm.transfer thy;
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1429
    val args = args_of thm;
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1430
    val incr_idx = Logic.incr_indexes ([], Thm.maxidx_of thm + 1);
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1431
    fun matches_args args' =
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1432
      let
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1433
        val k = length args - length args'
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1434
      in k >= 0 andalso Pattern.matchess thy (map incr_idx args', drop k args) end;
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1435
    fun matches (thm', proper') =
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1436
      (not proper orelse proper') andalso matches_args (args_of thm');
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1437
  in
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1438
    if exists matches eqns then
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1439
      (if verbose then warning ("Code generator: ignoring syntactically subsumed code equation\n" ^
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1440
        Thm.string_of_thm_global thy thm) else ();
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1441
      eqns)
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1442
    else
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1443
      eqns @ [(thm |> Thm.close_derivation \<^here> |> Thm.trim_context, proper)]
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1444
  end;
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1445
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1446
fun subsumptive_prepend thy { verbose } (thm, proper) eqns =
52637
1501ebe39711 attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents: 52475
diff changeset
  1447
  let
67522
9e712280cc37 clarified take/drop/chop prefix/suffix;
wenzelm
parents: 67147
diff changeset
  1448
    val args_of = drop_prefix is_Var o rev o snd o strip_comb
79411
700d4f16b5f2 minor performance tuning: proper Same.operation;
wenzelm
parents: 79232
diff changeset
  1449
      o Term.strip_sorts o fst o Logic.dest_equals o Thm.plain_prop_of
67646
85582dded912 trim context of persistent data;
wenzelm
parents: 67522
diff changeset
  1450
      o Thm.transfer thy;
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1451
    val args = args_of thm;
79232
99bc2dd45111 more general Logic.incr_indexes_operation;
wenzelm
parents: 78140
diff changeset
  1452
    val incr_idx = Logic.incr_indexes ([], Thm.maxidx_of thm + 1);
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1453
    fun matches_args args' =
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1454
      let
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1455
        val k = length args' - length args
82586
haftmann
parents: 81521
diff changeset
  1456
      in k >= 0 andalso Pattern.matchess thy (args, (map incr_idx o drop k) args') end;
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1457
    fun drop (thm', proper') = if (proper orelse not proper')
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1458
      andalso matches_args (args_of thm') then
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1459
        (if verbose then warning ("Code generator: dropping syntactically subsumed code equation\n" ^
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1460
          Thm.string_of_thm_global thy thm') else (); true)
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1461
      else false;
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 70358
diff changeset
  1462
  in (thm |> Thm.close_derivation \<^here> |> Thm.trim_context, proper) :: filter_out drop eqns end;
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
  1463
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1464
fun subsumptive_add thy { verbose, prepend } =
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1465
  if prepend then
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1466
    if Config.get_global thy prepend_allowed
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1467
    then subsumptive_prepend thy { verbose = verbose }
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1468
    else error "Not allowed to prepend code equation"
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1469
  else
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1470
    subsumptive_append thy { verbose = verbose };
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1471
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1472
fun add_eqn_for { check_singleton, prepend } (c, eqn) thy =
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1473
  thy |> (modify_specs o modify_pending_eqns thy { check_singleton = check_singleton } c)
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1474
    (subsumptive_add thy { verbose = true, prepend = prepend } eqn);
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1475
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1476
fun add_eqns_for { default } (c, proto_eqns) thy =
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1477
  thy |> modify_specs (fn specs =>
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1478
    if is_default (lookup_fun_spec specs c) orelse not default
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1479
    then
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1480
      let
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1481
        val eqns = []
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1482
          |> fold (subsumptive_append thy { verbose = not default }) proto_eqns;
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1483
      in specs |> register_fun_spec c (Eqns (default, eqns)) end
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1484
    else specs);
63242
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
  1485
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1486
fun add_abstract_for (c, (thm, tyco_abs as (tyco, _))) =
78056
904242f46ce1 more accurate Thm.trim_context / Thm.transfer;
wenzelm
parents: 78054
diff changeset
  1487
  modify_specs (register_fun_spec c (Abstr (Thm.close_derivation \<^here> thm |> Thm.trim_context, tyco_abs))
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1488
    #> map_types (History.modify_entry tyco (add_abstract_function c)))
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1489
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1490
in
63242
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
  1491
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1492
fun generic_declare_eqns { default } strictness raw_eqns thy =
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1493
  fold (add_eqns_for { default = default }) (prep_eqns strictness thy raw_eqns) thy;
37425
b5492f611129 explicitly name and note equations for class eq
haftmann
parents: 36615
diff changeset
  1494
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1495
fun generic_add_eqn { strictness, prepend } raw_eqn thy =
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1496
  fold (add_eqn_for { check_singleton = false, prepend = prepend }) (the_list (prep_eqn strictness thy raw_eqn)) thy;
52637
1501ebe39711 attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents: 52475
diff changeset
  1497
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1498
fun generic_declare_abstract_eqn strictness raw_abs_eqn thy =
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1499
  fold add_abstract_for (the_list (prep_abs_eqn strictness thy raw_abs_eqn)) thy;
55720
f3a2931a6656 add missing attributes
kuncar
parents: 55364
diff changeset
  1500
63242
9986559617ee clear distinction between different situations concerning strictness of code equations
haftmann
parents: 63241
diff changeset
  1501
fun add_maybe_abs_eqn_liberal thm thy =
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1502
  case prep_maybe_abs_eqn thy thm
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1503
   of SOME (c, (eqn, NONE)) => add_eqn_for { check_singleton = true, prepend = false } (c, eqn) thy
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1504
    | SOME (c, ((thm, _), SOME tyco)) => add_abstract_for (c, (thm, tyco)) thy
52637
1501ebe39711 attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents: 52475
diff changeset
  1505
    | NONE => thy;
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1506
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1507
end;
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1508
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1509
val declare_default_eqns_global = generic_declare_eqns { default = true } Silent;
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1510
val declare_default_eqns = code_declaration Silent (map o apfst o Morphism.thm)
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1511
  (generic_declare_eqns { default = true });
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1512
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1513
val declare_eqns_global = generic_declare_eqns { default = false } Strict;
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1514
val declare_eqns = code_declaration Liberal (map o apfst o Morphism.thm)
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1515
  (generic_declare_eqns  { default = false });
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1516
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1517
val add_eqn_global = generic_add_eqn { strictness = Strict, prepend = false };
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1518
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1519
fun del_eqn_global thm thy =
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1520
  case prep_eqn Liberal thy (thm, false) of
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1521
    SOME (c, (thm, _)) =>
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1522
      (modify_specs o modify_pending_eqns thy { check_singleton = false } c)
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1523
        (filter_out (fn (thm', _) => Thm.eq_thm_prop (thm, thm'))) thy
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
  1524
  | NONE => thy;
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
  1525
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1526
val declare_abstract_eqn_global = generic_declare_abstract_eqn Strict;
78068
a01c3bcf22dd tuned: avoid duplication;
wenzelm
parents: 78056
diff changeset
  1527
val declare_abstract_eqn = code_declaration Liberal Morphism.thm generic_declare_abstract_eqn;
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 54888
diff changeset
  1528
67032
ed499d1252fc strip some trailing spaces to force Pure rebuild after ce6454669360
Lars Hupel <lars.hupel@mytum.de>
parents: 66332
diff changeset
  1529
fun declare_aborting_global c =
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1530
  modify_specs (register_fun_spec c aborting);
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1531
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1532
fun declare_unimplemented_global c =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1533
  modify_specs (register_fun_spec c unimplemented);
34244
03f8dcab55f3 code cache without copy; tuned
haftmann
parents: 34184
diff changeset
  1534
03f8dcab55f3 code cache without copy; tuned
haftmann
parents: 34184
diff changeset
  1535
03f8dcab55f3 code cache without copy; tuned
haftmann
parents: 34184
diff changeset
  1536
(* cases *)
03f8dcab55f3 code cache without copy; tuned
haftmann
parents: 34184
diff changeset
  1537
40758
4f2c3e842ef8 consider sort constraints for datatype constructors when constructing the empty equation certificate;
haftmann
parents: 40726
diff changeset
  1538
fun case_cong thy case_const (num_args, (pos, _)) =
37438
4906ab970316 maintain cong rules for case combinators
haftmann
parents: 37425
diff changeset
  1539
  let
81521
1bfad73ab115 clarified signature: more operations;
wenzelm
parents: 81507
diff changeset
  1540
    val x :: y :: zs = Name.variants Name.context ("A" :: "A'" :: replicate (num_args - 1) "");
37438
4906ab970316 maintain cong rules for case combinators
haftmann
parents: 37425
diff changeset
  1541
    val (ws, vs) = chop pos zs;
54604
1512fa5fe531 prefer sort-stripping const_typ over Sign.the_const_type whenever appropriate
haftmann
parents: 54603
diff changeset
  1542
    val T = devarify (const_typ thy case_const);
40844
5895c525739d more direct use of binder_types/body_type;
wenzelm
parents: 40803
diff changeset
  1543
    val Ts = binder_types T;
37438
4906ab970316 maintain cong rules for case combinators
haftmann
parents: 37425
diff changeset
  1544
    val T_cong = nth Ts pos;
4906ab970316 maintain cong rules for case combinators
haftmann
parents: 37425
diff changeset
  1545
    fun mk_prem z = Free (z, T_cong);
4906ab970316 maintain cong rules for case combinators
haftmann
parents: 37425
diff changeset
  1546
    fun mk_concl z = list_comb (Const (case_const, T), map2 (curry Free) (ws @ z :: vs) Ts);
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58666
diff changeset
  1547
    val (prem, concl) = apply2 Logic.mk_equals (apply2 mk_prem (x, y), apply2 mk_concl (x, y));
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54604
diff changeset
  1548
  in
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54604
diff changeset
  1549
    Goal.prove_sorry_global thy (x :: y :: zs) [prem] concl
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54604
diff changeset
  1550
      (fn {context = ctxt', prems} =>
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54604
diff changeset
  1551
        Simplifier.rewrite_goals_tac ctxt' prems
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54604
diff changeset
  1552
        THEN ALLGOALS (Proof_Context.fact_tac ctxt' [Drule.reflexive_thm]))
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54604
diff changeset
  1553
  end;
37438
4906ab970316 maintain cong rules for case combinators
haftmann
parents: 37425
diff changeset
  1554
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1555
fun declare_case_global thm thy =
34244
03f8dcab55f3 code cache without copy; tuned
haftmann
parents: 34184
diff changeset
  1556
  let
43634
93cd6dd1a1c6 index cases for constructors
haftmann
parents: 43329
diff changeset
  1557
    val (case_const, (k, cos)) = case_cert thm;
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1558
    fun get_type_of_constr c = case get_type_of_constr_or_abstr thy c of
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1559
        SOME (c, false) => SOME c
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1560
      | _ => NONE;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1561
    val cos_with_tycos =
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1562
      (map_filter o Option.map) (fn c => (c, get_type_of_constr c)) cos;
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1563
    val _ = case map_filter (fn (c, NONE) => SOME c | _ => NONE) cos_with_tycos of
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1564
        [] => ()
45430
b8eb7a791dac misc tuning;
wenzelm
parents: 45344
diff changeset
  1565
      | cs => error ("Non-constructor(s) in case certificate: " ^ commas_quote cs);
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1566
    val tycos = distinct (op =) (map_filter snd cos_with_tycos);
75399
cdf84288d93c pass constructor arity as part of case certficiate
haftmann
parents: 75353
diff changeset
  1567
    val schema = (1 + Int.max (1, length cos),
cdf84288d93c pass constructor arity as part of case certficiate
haftmann
parents: 75353
diff changeset
  1568
      (k, (map o Option.map) (fn c => (c, args_number thy c)) cos));
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1569
    val cong = case_cong thy case_const schema;
37438
4906ab970316 maintain cong rules for case combinators
haftmann
parents: 37425
diff changeset
  1570
  in
4906ab970316 maintain cong rules for case combinators
haftmann
parents: 37425
diff changeset
  1571
    thy
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1572
    |> modify_specs (map_cases (History.register case_const
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1573
         (Case {schema = schema, tycos = tycos, cong = cong}))
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1574
      #> map_types (fold (fn tyco => History.modify_entry tyco
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1575
        (add_case_combinator case_const)) tycos))
37438
4906ab970316 maintain cong rules for case combinators
haftmann
parents: 37425
diff changeset
  1576
  end;
34244
03f8dcab55f3 code cache without copy; tuned
haftmann
parents: 34184
diff changeset
  1577
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1578
fun declare_undefined_global c =
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1579
  (modify_specs o map_cases) (History.register c Undefined);
34244
03f8dcab55f3 code cache without copy; tuned
haftmann
parents: 34184
diff changeset
  1580
03f8dcab55f3 code cache without copy; tuned
haftmann
parents: 34184
diff changeset
  1581
66310
e8d2862ec203 simplified function specification history: each pending function specification is historized at the end of a theory, without additional bookkeeping;
haftmann
parents: 66260
diff changeset
  1582
(* attributes *)
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31971
diff changeset
  1583
82858
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
  1584
fun code_attribute decl = Thm.declaration_attribute
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
  1585
  (fn thm => Context.mapping (decl thm) I);
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1586
82858
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
  1587
fun code_thm_attribute qualifier decl =
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
  1588
  Scan.lift (qualifier |-- Scan.succeed (code_attribute decl));
68773
haftmann
parents: 67649
diff changeset
  1589
82858
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
  1590
fun code_const_attribute qualifier decl =
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
  1591
  Scan.lift (qualifier -- Args.colon) |-- prospective_const_args
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82774
diff changeset
  1592
  >> (fn prospective_cs => code_attribute (fn _ => fn thy => fold decl (prospective_cs thy) thy));
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1593
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 52788
diff changeset
  1594
val _ = Theory.setup
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
  1595
  (let
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31971
diff changeset
  1596
    val code_attribute_parser =
68773
haftmann
parents: 67649
diff changeset
  1597
      code_thm_attribute (Args.$$$ "equation")
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1598
        (fn thm => generic_add_eqn { strictness = Liberal, prepend = false } (thm, true))
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1599
      || code_thm_attribute (Args.$$$ "prepend")
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1600
          (fn thm => generic_add_eqn { strictness = Liberal, prepend = true } (thm, true))
68773
haftmann
parents: 67649
diff changeset
  1601
      || code_thm_attribute (Args.$$$ "nbe")
82774
2865a6618cba append (rather than prepend) code equations: the order within a theory is maintained in the resulting code
haftmann
parents: 82773
diff changeset
  1602
          (fn thm => generic_add_eqn { strictness = Liberal, prepend = false } (thm, false))
68773
haftmann
parents: 67649
diff changeset
  1603
      || code_thm_attribute (Args.$$$ "abstract")
haftmann
parents: 67649
diff changeset
  1604
          (generic_declare_abstract_eqn Liberal)
haftmann
parents: 67649
diff changeset
  1605
      || code_thm_attribute (Args.$$$ "abstype")
haftmann
parents: 67649
diff changeset
  1606
          (generic_declare_abstype Liberal)
haftmann
parents: 67649
diff changeset
  1607
      || code_thm_attribute Args.del
haftmann
parents: 67649
diff changeset
  1608
          del_eqn_global
haftmann
parents: 67649
diff changeset
  1609
      || code_const_attribute (Args.$$$ "abort")
haftmann
parents: 67649
diff changeset
  1610
          declare_aborting_global
haftmann
parents: 67649
diff changeset
  1611
      || code_const_attribute (Args.$$$ "drop")
haftmann
parents: 67649
diff changeset
  1612
          declare_unimplemented_global
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1613
      || Scan.succeed (code_attribute
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 66189
diff changeset
  1614
          add_maybe_abs_eqn_liberal);
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
  1615
  in
73968
0274d442b7ea proper local context
haftmann
parents: 72057
diff changeset
  1616
    Attrib.setup \<^binding>\<open>code\<close> code_attribute_parser
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31971
diff changeset
  1617
        "declare theorems for code generation"
53171
a5e54d4d9081 added Theory.setup convenience;
wenzelm
parents: 52788
diff changeset
  1618
  end);
31962
baa8dce5bc45 tuned structure Code internally
haftmann
parents: 31957
diff changeset
  1619
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1620
end; (*struct*)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1621
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1622
35226
b987b803616d a simple concept for datatype invariants
haftmann
parents: 34901
diff changeset
  1623
(* type-safe interfaces for data dependent on executable code *)
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1624
34173
458ced35abb8 reduced code generator cache to the baremost minimum
haftmann
parents: 33977
diff changeset
  1625
functor Code_Data(Data: CODE_DATA_ARGS): CODE_DATA =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1626
struct
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1627
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1628
type T = Data.T;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1629
exception Data of T;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1630
fun dest (Data x) = x
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1631
34173
458ced35abb8 reduced code generator cache to the baremost minimum
haftmann
parents: 33977
diff changeset
  1632
val kind = Code.declare_data (Data Data.empty);
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1633
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1634
val data_op = (kind, Data, dest);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1635
39397
9b0a8d72edc8 ignore code cache optionally
haftmann
parents: 39134
diff changeset
  1636
fun change_yield (SOME thy) f = Code.change_yield_data data_op thy f
9b0a8d72edc8 ignore code cache optionally
haftmann
parents: 39134
diff changeset
  1637
  | change_yield NONE f = f Data.empty
9b0a8d72edc8 ignore code cache optionally
haftmann
parents: 39134
diff changeset
  1638
9b0a8d72edc8 ignore code cache optionally
haftmann
parents: 39134
diff changeset
  1639
fun change some_thy f = snd (change_yield some_thy (pair () o f));
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1640
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1641
end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
  1642
28143
e5c6c4aac52c different bookkeeping for code equations
haftmann
parents: 28054
diff changeset
  1643
structure Code : CODE = struct open Code; end;