src/HOL/Tools/record_package.ML
author wenzelm
Thu, 18 Jun 1998 18:35:07 +0200
changeset 5052 bbe3584b515b
parent 5006 cdc86a914e63
child 5060 7b86df67cc1a
permissions -rw-r--r--
fixed comment;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Tools/record_package.ML
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
     3
    Author:     Wolfgang Naraschewski and Markus Wenzel, TU Muenchen
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
     4
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
     5
Extensible records with structural subtyping in HOL.
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
     6
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
     7
TODO:
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
     8
  - field types: typedef;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
     9
  - trfuns for record types;
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    10
  - operations and theorems: split, split_all/ex, ...;
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
    11
  - field constructor: more specific type for snd component;
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    12
  - update_more operation;
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
    13
  - field syntax: "..." for "... = more", "?..." for "?... = ?more";
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    14
*)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    15
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    16
signature RECORD_PACKAGE =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    17
sig
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    18
  val moreS: sort
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    19
  val mk_fieldT: (string * typ) * typ -> typ
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    20
  val dest_fieldT: typ -> (string * typ) * typ
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    21
  val mk_field: (string * term) * term -> term
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    22
  val mk_fst: term -> term
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    23
  val mk_snd: term -> term
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    24
  val mk_recordT: (string * typ) list * typ -> typ
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    25
  val dest_recordT: typ -> (string * typ) list * typ
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    26
  val mk_record: (string * term) list * term -> term
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    27
  val mk_sel: term -> string -> term
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    28
  val mk_update: term -> string * term -> term
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    29
  val print_records: theory -> unit
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    30
  val add_record: (string list * bstring) -> string option
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    31
    -> (bstring * string) list -> theory -> theory
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    32
  val add_record_i: (string list * bstring) -> (typ list * string) option
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    33
    -> (bstring * typ) list -> theory -> theory
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    34
  val setup: (theory -> theory) list
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    35
end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    36
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    37
structure RecordPackage: RECORD_PACKAGE =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    38
struct
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    39
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    40
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    41
(*** utilities ***)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    42
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    43
(* string suffixes *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    44
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    45
fun suffix sfx s = s ^ sfx;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    46
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    47
fun unsuffix sfx s =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    48
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    49
    val cs = explode s;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    50
    val prfx_len = size s - size sfx;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    51
  in
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    52
    if prfx_len >= 0 andalso implode (drop (prfx_len, cs)) = sfx then
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    53
      implode (take (prfx_len, cs))
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    54
    else raise LIST "unsuffix"
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    55
  end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    56
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    57
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    58
(* definitions and equations *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    59
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    60
infix 0 :== === ;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    61
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    62
val (op :==) = Logic.mk_defpair;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    63
val (op ===) = HOLogic.mk_Trueprop o HOLogic.mk_eq;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    64
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    65
fun get_defs thy specs = map (PureThy.get_tthm thy o fst) specs;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    66
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    67
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    68
(* proof by simplification *)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    69
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
    70
fun prove_simp thy simps =
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    71
  let
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
    72
    val sign = Theory.sign_of thy;
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
    73
    val ss = Simplifier.addsimps (HOL_basic_ss, map Attribute.thm_of simps);
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
    74
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
    75
    fun prove goal =
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    76
      Attribute.tthm_of
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
    77
        (Goals.prove_goalw_cterm [] (Thm.cterm_of sign goal)
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    78
          (K [ALLGOALS (Simplifier.simp_tac ss)])
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    79
        handle ERROR => error ("The error(s) above occurred while trying to prove "
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
    80
          ^ quote (Sign.string_of_term sign goal)));
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    81
  in prove end;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    82
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    83
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    84
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    85
(*** syntax operations ***)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    86
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    87
(** name components **)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    88
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    89
val moreN = "more";
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    90
val schemeN = "_scheme";
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    91
val fieldN = "_field";
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    92
val field_typeN = "_field_type";
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    93
val fstN = "_fst";
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    94
val sndN = "_snd";
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    95
val updateN = "_update";
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    96
val makeN = "make";
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    97
val make_schemeN = "make_scheme";
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    98
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    99
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   100
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   101
(** tuple operations **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   102
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   103
(* more type class *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   104
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   105
val moreS = ["more"];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   106
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   107
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   108
(* types *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   109
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   110
fun mk_fieldT ((c, T), U) = Type (suffix field_typeN c, [T, U]);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   111
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   112
fun dest_fieldT (typ as Type (c_field_type, [T, U])) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   113
      (case try (unsuffix field_typeN) c_field_type of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   114
        None => raise TYPE ("dest_fieldT", [typ], [])
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   115
      | Some c => ((c, T), U))
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   116
  | dest_fieldT typ = raise TYPE ("dest_fieldT", [typ], []);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   117
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   118
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   119
(* constructors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   120
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   121
fun mk_fieldC U (c, T) = (suffix fieldN c, T --> U --> mk_fieldT ((c, T), U));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   122
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   123
fun mk_field ((c, t), u) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   124
  let val T = fastype_of t and U = fastype_of u
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   125
  in Const (suffix fieldN c, [T, U] ---> mk_fieldT ((c, T), U)) $ t $ u end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   126
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   127
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   128
(* destructors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   129
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   130
fun mk_fstC U (c, T) = (suffix fstN c, mk_fieldT ((c, T), U) --> T);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   131
fun mk_sndC U (c, T) = (suffix sndN c, mk_fieldT ((c, T), U) --> U);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   132
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   133
fun dest_field fst_or_snd p =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   134
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   135
    val pT = fastype_of p;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   136
    val ((c, T), U) = dest_fieldT pT;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   137
    val (destN, destT) = if fst_or_snd then (fstN, T) else (sndN, U);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   138
  in Const (suffix destN c, pT --> destT) $ p end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   139
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   140
val mk_fst = dest_field true;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   141
val mk_snd = dest_field false;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   142
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   143
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   144
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   145
(** record operations **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   146
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   147
(* types *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   148
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   149
val mk_recordT = foldr mk_fieldT;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   150
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   151
fun dest_recordT T =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   152
  (case try dest_fieldT T of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   153
    None => ([], T)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   154
  | Some (c_T, U) => apfst (cons c_T) (dest_recordT U));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   155
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   156
fun find_fieldT c rT =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   157
  (case assoc (fst (dest_recordT rT), c) of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   158
    None => raise TYPE ("find_field: " ^ c, [rT], [])
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   159
  | Some T => T);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   160
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   161
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   162
(* constructors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   163
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   164
val mk_record = foldr mk_field;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   165
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   166
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   167
(* selectors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   168
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   169
fun mk_selC rT (c, T) = (c, rT --> T);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   170
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   171
fun mk_sel r c =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   172
  let val rT = fastype_of r
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   173
  in Const (mk_selC rT (c, find_fieldT c rT)) $ r end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   174
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   175
val mk_moreC = mk_selC;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   176
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   177
fun mk_more r c =
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   178
  let val rT = fastype_of r
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   179
  in Const (mk_moreC rT (c, snd (dest_recordT rT))) $ r end;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   180
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   181
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   182
(* updates *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   183
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   184
fun mk_updateC rT (c, T) = (suffix updateN c, T --> rT --> rT);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   185
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   186
fun mk_update r (c, x) =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   187
  let val rT = fastype_of r
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   188
  in Const (mk_updateC rT (c, find_fieldT c rT)) $ x $ r end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   189
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   190
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   191
(* make *)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   192
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   193
fun mk_makeC rT (c, Ts) = (c, Ts ---> rT);
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   194
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   195
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   196
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   197
(** concrete syntax for records **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   198
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   199
(* parse translations *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   200
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   201
fun field_tr (Const ("_field", _) $ Free (name, _) $ arg) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   202
      Syntax.const (suffix fieldN name) $ arg
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   203
  | field_tr t = raise TERM ("field_tr", [t]);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   204
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   205
fun fields_tr (Const ("_fields", _) $ field $ fields) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   206
      field_tr field :: fields_tr fields
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   207
  | fields_tr field = [field_tr field];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   208
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   209
fun record_tr (*"_record"*) [fields] =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   210
      foldr (op $) (fields_tr fields, HOLogic.unit)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   211
  | record_tr (*"_record"*) ts = raise TERM ("record_tr", ts);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   212
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   213
fun record_scheme_tr (*"_record_scheme"*) [fields, more] =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   214
      foldr (op $) (fields_tr fields, more)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   215
  | record_scheme_tr (*"_record_scheme"*) ts = raise TERM ("record_scheme_tr", ts);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   216
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   217
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   218
(* print translations *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   219
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   220
fun fields_tr' (tm as Const (name_field, _) $ arg $ more) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   221
      (case try (unsuffix fieldN) name_field of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   222
        Some name =>
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   223
          apfst (cons (Syntax.const "_field" $ Syntax.free name $ arg)) (fields_tr' more)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   224
      | None => ([], tm))
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   225
  | fields_tr' tm = ([], tm);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   226
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   227
fun record_tr' tm =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   228
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   229
    val (fields, more) = fields_tr' tm;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   230
    val fields' = foldr1 (fn (f, fs) => Syntax.const "_fields" $ f $ fs) fields;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   231
  in
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   232
    if HOLogic.is_unit more then Syntax.const "_record" $ fields'
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   233
    else Syntax.const "_record_scheme" $ fields' $ more
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   234
  end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   235
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   236
fun field_tr' name [arg, more] = record_tr' (Syntax.const name $ arg $ more)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   237
  | field_tr' _ _ = raise Match;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   238
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   239
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   240
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   241
(*** extend theory by record definition ***)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   242
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   243
(** record info **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   244
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   245
(* type record_info and parent_info *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   246
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   247
type record_info =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   248
 {args: (string * sort) list,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   249
  parent: (typ list * string) option,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   250
  fields: (string * typ) list,
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   251
  simps: tthm list};
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   252
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   253
type parent_info =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   254
 {name: string,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   255
  fields: (string * typ) list,
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   256
  simps: tthm list};
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   257
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   258
5052
bbe3584b515b fixed comment;
wenzelm
parents: 5006
diff changeset
   259
(* data kind 'HOL/records' *)
5001
9de7fda0a6df accomodate tuned version of theory data;
wenzelm
parents: 4970
diff changeset
   260
5006
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   261
structure RecordsArgs =
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   262
struct
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   263
  val name = "HOL/records";
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   264
  type T = record_info Symtab.table;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   265
5006
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   266
  val empty = Symtab.empty;
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   267
  val prep_ext = I;
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   268
  val merge: T * T -> T = Symtab.merge (K true);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   269
5006
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   270
  fun print sg tab =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   271
    let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   272
      val prt_typ = Sign.pretty_typ sg;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   273
      val ext_const = Sign.cond_extern sg Sign.constK;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   274
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   275
      fun pretty_parent None = []
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   276
        | pretty_parent (Some (Ts, name)) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   277
            [Pretty.block [prt_typ (Type (name, Ts)), Pretty.str " +"]];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   278
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   279
      fun pretty_field (c, T) = Pretty.block
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   280
        [Pretty.str (ext_const c), Pretty.str " ::", Pretty.brk 1, Pretty.quote (prt_typ T)];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   281
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   282
      fun pretty_record (name, {args, parent, fields, simps = _}) = Pretty.block (Pretty.fbreaks
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   283
        (Pretty.block [prt_typ (Type (name, map TFree args)), Pretty.str " = "] ::
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   284
          pretty_parent parent @ map pretty_field fields));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   285
    in
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   286
      seq (Pretty.writeln o pretty_record) (Symtab.dest tab)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   287
    end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   288
end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   289
5006
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   290
structure RecordsData = TheoryDataFun(RecordsArgs);
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   291
val print_records = RecordsData.print;
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   292
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   293
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   294
(* get and put records *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   295
5006
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   296
fun get_record thy name = Symtab.lookup (RecordsData.get thy, name);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   297
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   298
fun put_record name info thy =
5006
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   299
  RecordsData.put (Symtab.update ((name, info), RecordsData.get thy)) thy;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   300
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   301
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   302
(* parent records *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   303
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   304
fun inst_record thy (types, name) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   305
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   306
    val sign = Theory.sign_of thy;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   307
    fun err msg = error (msg ^ " parent record " ^ quote name);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   308
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   309
    val {args, parent, fields, simps} =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   310
      (case get_record thy name of Some info => info | None => err "Unknown");
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   311
    val _ = if length types <> length args then err "Bad number of arguments for" else ();
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   312
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   313
    fun bad_inst ((x, S), T) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   314
      if Sign.of_sort sign (T, S) then None else Some x
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   315
    val bads = mapfilter bad_inst (args ~~ types);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   316
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   317
    val inst = map fst args ~~ types;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   318
    val subst = Term.map_type_tfree (fn (x, _) => the (assoc (inst, x)));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   319
  in
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   320
    if not (null bads) then
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   321
      err ("Ill-sorted instantiation of " ^ commas bads ^ " in")
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   322
    else (apsome (apfst (map subst)) parent, map (apsnd subst) fields, simps)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   323
  end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   324
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   325
fun add_parents thy (None, parents) = parents
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   326
  | add_parents thy (Some (types, name), parents) =
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   327
      let val (pparent, pfields, psimps) = inst_record thy (types, name)
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   328
      in add_parents thy (pparent, {name = name, fields = pfields, simps = psimps} :: parents) end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   329
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   330
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   331
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   332
(** internal theory extenders **)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   333
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   334
(* field_definitions *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   335
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   336
(*theorems from Prod.thy*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   337
val prod_convs = map Attribute.tthm_of [fst_conv, snd_conv];
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   338
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   339
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   340
fun field_definitions fields names zeta moreT more vars named_vars thy =
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   341
  let
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   342
    val base = Sign.base_name;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   343
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   344
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   345
    (* prepare declarations and definitions *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   346
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   347
    (*field types*)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   348
    fun mk_fieldT_spec c =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   349
      (suffix field_typeN c, ["'a", zeta],
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   350
        HOLogic.mk_prodT (TFree ("'a", HOLogic.termS), moreT), Syntax.NoSyn);
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   351
    val fieldT_specs = map (mk_fieldT_spec o base) names;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   352
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   353
    (*field declarations*)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   354
    val field_decls = map (mk_fieldC moreT) fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   355
    val dest_decls = map (mk_fstC moreT) fields @ map (mk_sndC moreT) fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   356
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   357
    (*field constructors*)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   358
    fun mk_field_spec (c, v) =
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   359
      mk_field ((c, v), more) :== HOLogic.mk_prod (v, more);
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   360
    val field_specs = map mk_field_spec named_vars;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   361
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   362
    (*field destructors*)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   363
    fun mk_dest_spec dest dest' (c, T) =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   364
      let
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   365
        val p = Free ("p", mk_fieldT ((c, T), moreT));
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   366
        val p' = Free ("p", HOLogic.mk_prodT (T, moreT));
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   367
          (*note: field types are just abbreviations*)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   368
      in dest p :== dest' p' end;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   369
    val dest_specs =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   370
      map (mk_dest_spec mk_fst HOLogic.mk_fst) fields @
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   371
      map (mk_dest_spec mk_snd HOLogic.mk_snd) fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   372
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   373
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   374
    (* prepare theorems *)
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   375
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   376
    fun mk_dest_prop dest dest' (c, v) =
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   377
      dest (mk_field ((c, v), more)) === dest' (v, more);
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   378
    val dest_props =
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   379
      map (mk_dest_prop mk_fst fst) named_vars @
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   380
      map (mk_dest_prop mk_snd snd) named_vars;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   381
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   382
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   383
    (* 1st stage: defs_thy *)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   384
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   385
    val defs_thy =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   386
      thy
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   387
      |> Theory.add_tyabbrs_i fieldT_specs
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   388
      |> (Theory.add_consts_i o map (Syntax.no_syn o apfst base))
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   389
        (field_decls @ dest_decls)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   390
      |> (PureThy.add_defs_i o map Attribute.none)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   391
        (field_specs @ dest_specs);
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   392
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   393
    val field_defs = get_defs defs_thy field_specs;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   394
    val dest_defs = get_defs defs_thy dest_specs;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   395
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   396
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   397
    (* 2nd stage: thms_thy *)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   398
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   399
    val dest_convs =
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   400
      map (prove_simp defs_thy (field_defs @ dest_defs @ prod_convs)) dest_props;
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   401
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   402
    val thms_thy =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   403
      defs_thy
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   404
      |> (PureThy.add_tthmss o map Attribute.none)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   405
        [("field_defs", field_defs),
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   406
          ("dest_defs", dest_defs),
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   407
          ("dest_convs", dest_convs)];
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   408
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   409
  in (thms_thy, dest_convs) end;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   410
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   411
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   412
(* record_definition *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   413
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   414
fun record_definition (args, bname) parent (parents: parent_info list) bfields thy =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   415
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   416
    val sign = Theory.sign_of thy;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   417
    val full = Sign.full_name_path sign bname;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   418
    val base = Sign.base_name;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   419
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   420
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   421
    (* basic components *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   422
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   423
    val alphas = map fst args;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   424
    val name = Sign.full_name sign bname;	(*not made part of record name space!*)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   425
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   426
    val parent_fields = flat (map #fields parents);
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   427
    val parent_names = map fst parent_fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   428
    val parent_types = map snd parent_fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   429
    val parent_len = length parent_fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   430
    val parent_xs = variantlist (map (base o fst) parent_fields, [moreN]);
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   431
    val parent_vars = ListPair.map Free (parent_xs, parent_types);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   432
    val parent_named_vars = parent_names ~~ parent_vars;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   433
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   434
    val fields = map (apfst full) bfields;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   435
    val names = map fst fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   436
    val types = map snd fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   437
    val len = length fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   438
    val xs = variantlist (map fst bfields, moreN :: parent_xs);
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   439
    val vars = ListPair.map Free (xs, types);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   440
    val named_vars = names ~~ vars;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   441
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   442
    val all_fields = parent_fields @ fields;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   443
    val all_names = parent_names @ names;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   444
    val all_types = parent_types @ types;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   445
    val all_len = parent_len + len;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   446
    val all_xs = parent_xs @ xs;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   447
    val all_vars = parent_vars @ vars;
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   448
    val all_named_vars = parent_named_vars @ named_vars;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   449
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   450
    val zeta = variant alphas "'z";
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   451
    val moreT = TFree (zeta, moreS);
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   452
    val more = Free (moreN, moreT);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   453
    fun more_part t = mk_more t (full moreN);
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   454
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   455
    val parent_more = funpow parent_len mk_snd;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   456
    val idxs = 0 upto (len - 1);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   457
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   458
    val rec_schemeT = mk_recordT (all_fields, moreT);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   459
    val rec_scheme = mk_record (all_named_vars, more);
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   460
    val r = Free ("r", rec_schemeT);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   461
    val recT = mk_recordT (all_fields, HOLogic.unitT);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   462
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   463
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   464
    (* prepare print translation functions *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   465
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   466
    val field_tr'_names =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   467
      distinct (flat (map (NameSpace.accesses o suffix fieldN) names)) \\
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   468
        #3 (Syntax.trfun_names (Theory.syn_of thy));
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   469
    val field_trfuns = ([], [], field_tr'_names ~~ map field_tr' field_tr'_names, []);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   470
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   471
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   472
    (* prepare declarations *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   473
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   474
    val sel_decls = map (mk_selC rec_schemeT) bfields @ [mk_moreC rec_schemeT (moreN, moreT)];
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   475
    val update_decls = map (mk_updateC rec_schemeT) bfields;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   476
    val make_decls =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   477
      [(mk_makeC rec_schemeT (make_schemeN, all_types @ [moreT])),
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   478
       (mk_makeC recT (makeN, all_types))];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   479
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   480
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   481
    (* prepare definitions *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   482
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   483
    (*record (scheme) type abbreviation*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   484
    val recordT_specs =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   485
      [(suffix schemeN bname, alphas @ [zeta], rec_schemeT, Syntax.NoSyn),
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   486
        (bname, alphas, recT, Syntax.NoSyn)];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   487
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   488
    (*selectors*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   489
    fun mk_sel_spec (i, c) =
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   490
      mk_sel r c :== mk_fst (funpow i mk_snd (parent_more r));
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   491
    val sel_specs =
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   492
      ListPair.map mk_sel_spec (idxs, names) @
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   493
        [more_part r :== funpow len mk_snd (parent_more r)];
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   494
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   495
    (*updates*)
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   496
    val all_sels = all_names ~~ map (mk_sel r) all_names;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   497
    fun mk_upd_spec (i, (c, x)) =
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   498
      mk_update r (c, x) :==
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   499
        mk_record (nth_update (c, x) (parent_len + i, all_sels), more_part r)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   500
    val update_specs = ListPair.map mk_upd_spec (idxs, named_vars);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   501
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   502
    (*makes*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   503
    val make_scheme = Const (mk_makeC rec_schemeT (full make_schemeN, all_types @ [moreT]));
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   504
    val make = Const (mk_makeC recT (full makeN, all_types));
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   505
    val make_specs =
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   506
      [list_comb (make_scheme, all_vars) $ more :== rec_scheme,
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   507
        list_comb (make, all_vars) :== mk_record (all_named_vars, HOLogic.unit)];
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   508
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   509
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   510
    (* prepare propositions *)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   511
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   512
    (*selectors*)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   513
    val sel_props =
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   514
      map (fn (c, x) => mk_sel rec_scheme c === x) named_vars @
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   515
        [more_part rec_scheme === more];
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   516
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   517
    (*updates*)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   518
    fun mk_upd_prop (i, (c, T)) =
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   519
      let val x' = Free (variant all_xs (base c ^ "'"), T) in
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   520
        mk_update rec_scheme (c, x') ===
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   521
          mk_record (nth_update (c, x') (parent_len + i, all_named_vars), more)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   522
      end;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   523
    val update_props = ListPair.map mk_upd_prop (idxs, fields);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   524
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   525
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   526
    (* 1st stage: fields_thy *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   527
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   528
    val (fields_thy, field_simps) =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   529
      thy
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   530
      |> Theory.add_path bname
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   531
      |> field_definitions fields names zeta moreT more vars named_vars;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   532
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   533
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   534
    (* 2nd stage: defs_thy *)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   535
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   536
    val defs_thy =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   537
      fields_thy
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   538
      |> Theory.parent_path
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   539
      |> Theory.add_tyabbrs_i recordT_specs	(*not made part of record name space!*)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   540
      |> Theory.add_path bname
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   541
      |> Theory.add_trfuns field_trfuns
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   542
      |> (Theory.add_consts_i o map Syntax.no_syn)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   543
        (sel_decls @ update_decls @ make_decls)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   544
      |> (PureThy.add_defs_i o map Attribute.none)
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   545
        (sel_specs @ update_specs @ make_specs);
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   546
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   547
    val sel_defs = get_defs defs_thy sel_specs;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   548
    val update_defs = get_defs defs_thy update_specs;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   549
    val make_defs = get_defs defs_thy make_specs;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   550
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   551
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   552
    (* 3rd stage: thms_thy *)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   553
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   554
    val parent_simps = flat (map #simps parents);
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   555
    val prove = prove_simp defs_thy;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   556
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   557
    val sel_convs = map (prove (parent_simps @ sel_defs @ field_simps)) sel_props;
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   558
    val update_convs = map (prove (parent_simps @ update_defs @ sel_convs)) update_props;
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   559
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   560
    val simps = field_simps @ sel_convs @ update_convs @ make_defs;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   561
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   562
    val thms_thy =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   563
      defs_thy
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   564
      |> (PureThy.add_tthmss o map Attribute.none)
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   565
        [("select_defs", sel_defs),
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   566
          ("update_defs", update_defs),
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   567
          ("make_defs", make_defs),
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   568
          ("select_convs", sel_convs),
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   569
          ("update_convs", update_convs)]
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   570
      |> PureThy.add_tthmss [(("simps", simps), [Simplifier.simp_add_global])];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   571
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   572
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   573
    (* 4th stage: final_thy *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   574
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   575
    val final_thy =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   576
      thms_thy
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   577
      |> put_record name {args = args, parent = parent, fields = fields, simps = simps}
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   578
      |> Theory.parent_path;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   579
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   580
  in final_thy end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   581
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   582
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   583
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   584
(** theory extender interface **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   585
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   586
(* prepare arguments *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   587
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   588
(*note: read_raw_typ avoids expanding type abbreviations*)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   589
fun read_raw_parent sign s =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   590
  (case Sign.read_raw_typ (sign, K None) s handle TYPE (msg, _, _) => error msg of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   591
    Type (name, Ts) => (Ts, name)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   592
  | _ => error ("Bad parent record specification: " ^ quote s));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   593
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   594
fun read_typ sign (env, s) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   595
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   596
    fun def_type (x, ~1) = assoc (env, x)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   597
      | def_type _ = None;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   598
    val T = Type.no_tvars (Sign.read_typ (sign, def_type) s) handle TYPE (msg, _, _) => error msg;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   599
  in (Term.add_typ_tfrees (T, env), T) end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   600
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   601
fun cert_typ sign (env, raw_T) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   602
  let val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle TYPE (msg, _, _) => error msg
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   603
  in (Term.add_typ_tfrees (T, env), T) end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   604
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   605
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   606
(* add_record *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   607
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   608
(*we do all preparations and error checks here, deferring the real
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   609
  work to record_definition*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   610
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   611
fun gen_add_record prep_typ prep_raw_parent (params, bname) raw_parent raw_fields thy =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   612
  let
4970
8b65444edbb0 Changed require to requires for MLWorks
paulson
parents: 4967
diff changeset
   613
    val _ = Theory.requires thy "Record" "record definitions";
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   614
    val sign = Theory.sign_of thy;
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   615
    val _ = writeln ("Defining record " ^ quote bname ^ " ...");
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   616
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   617
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   618
    (* parents *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   619
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   620
    fun prep_inst T = snd (cert_typ sign ([], T));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   621
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   622
    val parent = apsome (apfst (map prep_inst) o prep_raw_parent sign) raw_parent
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   623
      handle ERROR => error ("The error(s) above in parent record specification");
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   624
    val parents = add_parents thy (parent, []);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   625
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   626
    val init_env =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   627
      (case parent of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   628
        None => []
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   629
      | Some (types, _) => foldr Term.add_typ_tfrees (types, []));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   630
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   631
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   632
    (* fields *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   633
4967
c9c481bc0216 foldl_map prep_field;
wenzelm
parents: 4895
diff changeset
   634
    fun prep_field (env, (c, raw_T)) =
c9c481bc0216 foldl_map prep_field;
wenzelm
parents: 4895
diff changeset
   635
      let val (env', T) = prep_typ sign (env, raw_T) handle ERROR =>
c9c481bc0216 foldl_map prep_field;
wenzelm
parents: 4895
diff changeset
   636
        error ("The error(s) above occured in field " ^ quote c)
c9c481bc0216 foldl_map prep_field;
wenzelm
parents: 4895
diff changeset
   637
      in (env', (c, T)) end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   638
4967
c9c481bc0216 foldl_map prep_field;
wenzelm
parents: 4895
diff changeset
   639
    val (envir, bfields) = foldl_map prep_field (init_env, raw_fields);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   640
    val envir_names = map fst envir;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   641
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   642
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   643
    (* args *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   644
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   645
    val defaultS = Sign.defaultS sign;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   646
    val args = map (fn x => (x, if_none (assoc (envir, x)) defaultS)) params;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   647
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   648
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   649
    (* errors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   650
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   651
    val name = Sign.full_name sign bname;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   652
    val err_dup_record =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   653
      if is_none (get_record thy name) then []
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   654
      else ["Duplicate definition of record " ^ quote name];
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   655
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   656
    val err_dup_parms =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   657
      (case duplicates params of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   658
        [] => []
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   659
      | dups => ["Duplicate parameter(s) " ^ commas dups]);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   660
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   661
    val err_extra_frees =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   662
      (case gen_rems (op =) (envir_names, params) of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   663
        [] => []
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   664
      | extras => ["Extra free type variable(s) " ^ commas extras]);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   665
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   666
    val err_no_fields = if null bfields then ["No fields present"] else [];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   667
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   668
    val err_dup_fields =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   669
      (case duplicates (map fst bfields) of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   670
        [] => []
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   671
      | dups => ["Duplicate field(s) " ^ commas_quote dups]);
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   672
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   673
    val err_bad_fields =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   674
      if forall (not_equal moreN o fst) bfields then []
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   675
      else ["Illegal field name " ^ quote moreN];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   676
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   677
    val err_dup_sorts =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   678
      (case duplicates envir_names of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   679
        [] => []
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   680
      | dups => ["Inconsistent sort constraints for " ^ commas dups]);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   681
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   682
    val errs =
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   683
      err_dup_record @ err_dup_parms @ err_extra_frees @ err_no_fields @
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   684
      err_dup_fields @ err_bad_fields @ err_dup_sorts;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   685
  in
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   686
    if null errs then () else error (cat_lines errs);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   687
    thy |> record_definition (args, bname) parent parents bfields
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   688
  end
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   689
  handle ERROR => error ("Failed to define record " ^ quote bname);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   690
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   691
val add_record = gen_add_record read_typ read_raw_parent;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   692
val add_record_i = gen_add_record cert_typ (K I);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   693
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   694
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   695
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   696
(** setup theory **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   697
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   698
val setup =
5006
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   699
 [RecordsData.init,
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   700
  Theory.add_trfuns
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   701
    ([], [("_record", record_tr), ("_record_scheme", record_scheme_tr)], [], [])];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   702
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   703
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   704
end;