src/HOL/Tools/record_package.ML
author wenzelm
Wed, 29 Apr 1998 11:40:37 +0200
changeset 4867 9be2bf0ce909
child 4890 f0a24bad990a
permissions -rw-r--r--
package extensible records with structural subtyping in HOL -- still experimental version;
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:
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
     8
  - record_info: tr' funs;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
     9
  - trfuns for record types;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    10
  - field types: typedef;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    11
  - make selector types as general as possible (no!?);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    12
*)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    13
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    14
signature RECORD_PACKAGE =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    15
sig
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    16
  val print_records: theory -> unit
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    17
  val add_record: (string list * bstring) -> string option
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    18
    -> (bstring * string) list -> theory -> theory
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    19
  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
    20
    -> (bstring * typ) list -> theory -> theory
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    21
  val setup: (theory -> theory) list
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    22
end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    23
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    24
structure RecordPackage: RECORD_PACKAGE =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    25
struct
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    26
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    27
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    28
(*** syntax operations ***)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    29
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    30
(** names **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    31
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    32
(* name components *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    33
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    34
val moreN = "more";
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    35
val schemeN = "_scheme";
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    36
val fieldN = "_field";
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    37
val field_typeN = "_field_type";
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    38
val fstN = "_val";
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    39
val sndN = "_more";
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    40
val updateN = "_update";
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    41
val makeN = "make";
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    42
val make_schemeN = "make_scheme";
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    43
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
(* suffixes *)
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 suffix sfx s = s ^ sfx;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    48
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    49
fun unsuffix sfx s =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    50
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    51
    val cs = explode s;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    52
    val prfx_len = size s - size sfx;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    53
  in
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    54
    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
    55
      implode (take (prfx_len, cs))
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    56
    else raise LIST "unsuffix"
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    57
  end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    58
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    59
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    60
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    61
(** tuple operations **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    62
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    63
(* more type class *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    64
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    65
val moreS = ["more"];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    66
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    67
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    68
(* types *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    69
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    70
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
    71
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    72
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
    73
      (case try (unsuffix field_typeN) c_field_type of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    74
        None => raise TYPE ("dest_fieldT", [typ], [])
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    75
      | Some c => ((c, T), U))
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    76
  | dest_fieldT typ = raise TYPE ("dest_fieldT", [typ], []);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    77
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    78
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    79
(* constructors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    80
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    81
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
    82
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    83
fun mk_field ((c, t), u) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    84
  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
    85
  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
    86
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    87
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    88
(* destructors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    89
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    90
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
    91
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
    92
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    93
fun dest_field fst_or_snd p =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    94
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    95
    val pT = fastype_of p;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    96
    val ((c, T), U) = dest_fieldT pT;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    97
    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
    98
  in Const (suffix destN c, pT --> destT) $ p end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    99
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   100
val mk_fst = dest_field true;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   101
val mk_snd = dest_field false;
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
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
(** record operations **)
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
(* types *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   108
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   109
val mk_recordT = foldr mk_fieldT;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   110
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   111
fun dest_recordT T =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   112
  (case try dest_fieldT T of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   113
    None => ([], T)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   114
  | 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
   115
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   116
fun find_fieldT c rT =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   117
  (case assoc (fst (dest_recordT rT), c) of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   118
    None => raise TYPE ("find_field: " ^ c, [rT], [])
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   119
  | Some T => T);
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
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   122
(* constructors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   123
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   124
val mk_record = foldr mk_field;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   125
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
(* selectors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   128
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   129
fun mk_selC rT (c, T) = (c, rT --> T);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   130
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   131
fun mk_sel c r =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   132
  let val rT = fastype_of r
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   133
  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
   134
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   135
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   136
(* updates *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   137
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   138
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
   139
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   140
fun mk_update c x r =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   141
  let val rT = fastype_of r
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   142
  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
   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
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   146
(** concrete syntax for records **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   147
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   148
(* parse translations *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   149
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   150
fun field_tr (Const ("_field", _) $ Free (name, _) $ arg) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   151
      Syntax.const (suffix fieldN name) $ arg
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   152
  | field_tr t = raise TERM ("field_tr", [t]);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   153
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   154
fun fields_tr (Const ("_fields", _) $ field $ fields) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   155
      field_tr field :: fields_tr fields
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   156
  | fields_tr field = [field_tr field];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   157
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   158
fun record_tr (*"_record"*) [fields] =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   159
      foldr (op $) (fields_tr fields, HOLogic.unit)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   160
  | record_tr (*"_record"*) ts = raise TERM ("record_tr", ts);
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
fun record_scheme_tr (*"_record_scheme"*) [fields, more] =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   163
      foldr (op $) (fields_tr fields, more)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   164
  | 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
   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
(* print translations *)		(* FIXME tune, activate *)
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
(* FIXME ... :: tms *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   170
fun fields_tr' (tm as Const (name_field, _) $ arg $ more) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   171
      (case try (unsuffix fieldN) name_field of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   172
        Some name =>
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   173
          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
   174
      | None => ([], tm))
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   175
  | fields_tr' tm = ([], tm);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   176
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   177
fun record_tr' tm =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   178
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   179
    val mk_fields = foldr (fn (field, fields) => Syntax.const "_fields" $ field $ fields);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   180
    val (fields, more) = fields_tr' tm;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   181
  in
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   182
    if HOLogic.is_unit more then
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   183
      Syntax.const "_record" $ mk_fields (split_last fields)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   184
    else Syntax.const "_record_scheme" $ mk_fields (fields, more)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   185
  end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   186
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   187
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
   188
  | field_tr' _ _ = raise Match;
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
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   191
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   192
(*** extend theory by record definition ***)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   193
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   194
(** record info **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   195
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   196
(* type record_info and parent_info *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   197
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   198
type record_info =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   199
 {args: (string * sort) list,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   200
  parent: (typ list * string) option,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   201
  fields: (string * typ) list,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   202
  simps: tthm list};
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   203
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   204
type parent_info =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   205
 {name: string,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   206
  fields: (string * typ) list,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   207
  simps: tthm list};
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
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   210
(* theory data *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   211
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   212
val recordsK = "HOL/records";
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   213
exception Records of record_info Symtab.table;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   214
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   215
fun print_records thy = Display.print_data thy recordsK;
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
local
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   218
  val empty = Records Symtab.empty;
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 prep_ext (x as Records _) = x;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   221
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   222
  fun merge (Records tab1, Records tab2) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   223
    Records (Symtab.merge (K true) (tab1, tab2));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   224
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   225
  fun print sg (Records tab) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   226
    let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   227
      val prt_typ = Sign.pretty_typ sg;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   228
      val ext_const = Sign.cond_extern sg Sign.constK;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   229
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   230
      fun pretty_parent None = []
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   231
        | pretty_parent (Some (Ts, name)) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   232
            [Pretty.block [prt_typ (Type (name, Ts)), Pretty.str " +"]];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   233
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   234
      fun pretty_field (c, T) = Pretty.block
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   235
        [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
   236
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   237
      fun pretty_record (name, {args, parent, fields, simps = _}) = Pretty.block (Pretty.fbreaks
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   238
        (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
   239
          pretty_parent parent @ map pretty_field fields));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   240
    in
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   241
      seq (Pretty.writeln o pretty_record) (Symtab.dest tab)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   242
    end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   243
in
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   244
  val record_thy_data = (recordsK, (empty, prep_ext, merge, print));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   245
end;
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
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   248
(* get and put records *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   249
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   250
fun get_records thy =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   251
  (case Theory.get_data thy recordsK of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   252
    Records tab => tab
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   253
  | _ => type_error recordsK);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   254
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   255
fun get_record thy name = Symtab.lookup (get_records thy, name);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   256
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
fun put_records tab thy =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   259
  Theory.put_data (recordsK, Records tab) thy;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   260
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   261
fun put_new_record name info thy =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   262
  thy |> put_records
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   263
    (Symtab.update_new ((name, info), get_records thy)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   264
      handle Symtab.DUP _ => error ("Duplicate definition of record " ^ quote name));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   265
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   266
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   267
(* parent records *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   268
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   269
fun inst_record thy (types, name) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   270
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   271
    val sign = Theory.sign_of thy;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   272
    fun err msg = error (msg ^ " parent record " ^ quote name);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   273
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   274
    val {args, parent, fields, simps} =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   275
      (case get_record thy name of Some info => info | None => err "Unknown");
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   276
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   277
    fun bad_inst ((x, S), T) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   278
      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
   279
    val bads = mapfilter bad_inst (args ~~ types);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   280
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   281
    val inst = map fst args ~~ types;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   282
    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
   283
  in
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   284
    if length types <> length args then
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   285
      err "Bad number of arguments for"
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   286
    else if not (null bads) then
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   287
      err ("Ill-sorted instantiation of " ^ commas bads ^ " in")
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   288
    else (apsome (apfst (map subst)) parent, map (apsnd subst) fields, simps)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   289
  end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   290
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   291
fun add_parents thy (None, parents) = parents
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   292
  | add_parents thy (Some (types, name), parents) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   293
      let val (pparent, pfields, psimps) = inst_record thy (types, name)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   294
      in add_parents thy (pparent, {name = name, fields = pfields, simps = psimps} :: parents) end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   295
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   296
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   297
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   298
(** record theorems **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   299
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   300
(* proof by simplification *)
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
fun prove_simp thy opt_ss simps =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   303
  let val ss = if_none opt_ss HOL_basic_ss addsimps simps in
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   304
    fn goal => Goals.prove_goalw_cterm [] (Thm.cterm_of (sign_of thy) goal)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   305
      (K [ALLGOALS (Simplifier.simp_tac ss)])
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   306
  end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   307
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   308
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   309
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   310
(** internal theory extender **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   311
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   312
(*do the actual record definition, assuming that all arguments are
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   313
  well-formed*)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   314
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   315
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
   316
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   317
    val sign = Theory.sign_of thy;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   318
    val full = Sign.full_name_path sign bname;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   319
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   320
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   321
    (* input *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   322
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   323
    val alphas = map fst args;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   324
    val name = Sign.full_name sign bname;		(* FIXME !? *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   325
    val parent_fields = flat (map #fields parents);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   326
    val fields = map (apfst full) bfields;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   327
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   328
    val all_fields = parent_fields @ fields;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   329
    val all_types = map snd all_fields;
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
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   332
    (* term / type components *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   333
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   334
    val zeta = variant alphas "'z";
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   335
    val moreT = TFree (zeta, moreS);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   336
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   337
    val xs = variantlist (map fst bfields, []);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   338
    val vars = map2 Free (xs, map snd fields);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   339
    val more = Free (variant xs moreN, moreT);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   340
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   341
    val rec_schemeT = mk_recordT (all_fields, moreT);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   342
    val recT = mk_recordT (all_fields, HOLogic.unitT);
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
    (* FIXME tune *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   345
    val make_schemeT = all_types ---> moreT --> rec_schemeT;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   346
    val make_scheme = Const (full make_schemeN, make_schemeT);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   347
    val makeT = all_types ---> recT;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   348
    val make = Const (full makeN, makeT);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   349
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   350
    val parent_more = funpow (length parent_fields) mk_snd;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   351
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   352
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   353
    (* prepare type definitions *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   354
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   355
    (*field types*)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   356
    fun mk_fieldT_spec ((c, T), a) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   357
      (suffix field_typeN c, [a, zeta],
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   358
        HOLogic.mk_prodT (TFree (a, HOLogic.termS), moreT), Syntax.NoSyn);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   359
    val fieldT_specs = map2 mk_fieldT_spec (bfields, alphas);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   360
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   361
    (*record types*)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   362
    val recordT_specs =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   363
      [(suffix schemeN bname, alphas @ [zeta], rec_schemeT, Syntax.NoSyn),
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   364
        (bname, alphas, recT, Syntax.NoSyn)];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   365
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   366
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   367
    (* prepare declarations *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   368
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   369
    val field_decls = map (mk_fieldC moreT) fields;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   370
    val dest_decls = map (mk_fstC moreT) fields @ map (mk_sndC moreT) fields;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   371
    val sel_decls = map (mk_selC rec_schemeT) fields;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   372
    val update_decls = map (mk_updateC rec_schemeT) fields;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   373
    val make_decls = [(make_schemeN, make_schemeT), (makeN, makeT)];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   374
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   375
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   376
    (* prepare definitions *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   377
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   378
    (*field constructors*)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   379
    fun mk_field_spec ((c, _), v) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   380
      Logic.mk_defpair (mk_field ((c, v), more), HOLogic.mk_prod (v, more));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   381
    val field_specs = map2 mk_field_spec (fields, vars);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   382
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   383
    (*field destructors*)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   384
    fun mk_dest_spec dest dest' (c, T) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   385
      let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   386
        val p = Free ("p",  mk_fieldT ((c, T), moreT));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   387
        val p' = Free ("p",  HOLogic.mk_prodT (T, moreT));  (*Note: field types are abbreviations*)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   388
      in Logic.mk_defpair (dest p, dest' p') end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   389
    val dest_specs =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   390
      map (mk_dest_spec mk_fst HOLogic.mk_fst) fields @
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   391
      map (mk_dest_spec mk_snd HOLogic.mk_snd) fields;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   392
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   393
    (*field selectors*)		(* FIXME tune *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   394
    fun mk_sel_specs _ [] specs = rev specs
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   395
      | mk_sel_specs prfx ((c, T) :: fs) specs =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   396
          let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   397
            val prfx' = prfx @ [(c, T)];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   398
            val r = Free ("r", mk_recordT (prfx' @ fs, moreT));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   399
            val spec = Logic.mk_defpair (mk_sel c r, mk_fst (funpow (length prfx) mk_snd r));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   400
          in mk_sel_specs prfx' fs (spec :: specs) end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   401
    val sel_specs = mk_sel_specs parent_fields fields [];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   402
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   403
    (*updates*)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   404
    val update_specs = [];	(* FIXME *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   405
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   406
    (*makes*)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   407
    val make_specs =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   408
      map Logic.mk_defpair
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   409
        [(list_comb (make_scheme, vars) $ more, mk_record (map fst fields ~~ vars, more)),
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   410
          (list_comb (make, vars), mk_record (map fst fields ~~ vars, HOLogic.unit))];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   411
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   412
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   413
    (* 1st stage: defs_thy *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   414
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   415
    val defs_thy =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   416
      thy
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   417
      |> Theory.add_path bname
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   418
      |> Theory.add_tyabbrs_i (fieldT_specs @ recordT_specs)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   419
      |> (Theory.add_consts_i o map (Syntax.no_syn o apfst Sign.base_name))
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   420
        (field_decls @ dest_decls @ sel_decls @ update_decls @ make_decls)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   421
      |> (PureThy.add_defs_i o map Attribute.none)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   422
        (field_specs @ dest_specs @ sel_specs @ update_specs @ make_specs);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   423
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   424
    local fun get_defs specs = map (PureThy.get_tthm defs_thy o fst) specs in
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   425
      val make_defs = get_defs make_specs;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   426
      val field_defs = get_defs field_specs;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   427
      val sel_defs = get_defs sel_specs;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   428
      val update_defs = get_defs update_specs;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   429
    end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   430
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   431
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   432
    (* 2nd stage: thms_thy *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   433
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   434
    val thms_thy =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   435
      defs_thy
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   436
      |> (PureThy.add_tthmss o map Attribute.none)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   437
        [("make_defs", make_defs),
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   438
          ("field_defs", field_defs),
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   439
          ("sel_defs", sel_defs),
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   440
          ("update_defs", update_defs)]
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   441
(*    |> record_theorems FIXME *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   442
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   443
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   444
    (* 3rd stage: final_thy *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   445
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   446
    val final_thy =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   447
      thms_thy
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   448
      |> put_new_record name
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   449
        {args = args, parent = parent, fields = fields, simps = [] (* FIXME *)}
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   450
      |> Theory.parent_path;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   451
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   452
  in final_thy end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   453
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   454
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   455
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   456
(** theory extender interface **)
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
(*do all preparations and error checks here, deferring the real work
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   459
  to record_definition above*)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   460
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   461
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   462
(* prepare arguments *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   463
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   464
(*Note: read_raw_typ avoids expanding type abbreviations*)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   465
fun read_raw_parent sign s =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   466
  (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
   467
    Type (name, Ts) => (Ts, name)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   468
  | _ => error ("Bad parent record specification: " ^ quote s));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   469
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   470
fun read_typ sign (env, s) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   471
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   472
    fun def_type (x, ~1) = assoc (env, x)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   473
      | def_type _ = None;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   474
    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
   475
  in (Term.add_typ_tfrees (T, env), T) end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   476
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   477
fun cert_typ sign (env, raw_T) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   478
  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
   479
  in (Term.add_typ_tfrees (T, env), T) end;
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
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   482
(* add_record *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   483
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   484
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
   485
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   486
    val _ = Theory.require thy "Record" "record definitions";
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   487
    val sign = Theory.sign_of thy;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   488
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   489
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   490
    (* parents *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   491
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   492
    fun prep_inst T = snd (cert_typ sign ([], T));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   493
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   494
    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
   495
      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
   496
    val parents = add_parents thy (parent, []);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   497
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   498
    val init_env =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   499
      (case parent of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   500
        None => []
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   501
      | Some (types, _) => foldr Term.add_typ_tfrees (types, []));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   502
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   503
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   504
    (* fields *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   505
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   506
    fun prep_fields (env, []) = (env, [])
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   507
      | prep_fields (env, (c, raw_T) :: fs) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   508
          let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   509
            val (env', T) = prep_typ sign (env, raw_T) handle ERROR =>
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   510
              error ("The error(s) above occured in field " ^ quote c);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   511
            val (env'', fs') = prep_fields (env', fs);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   512
      in (env'', (c, T) :: fs') end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   513
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   514
    val (envir, bfields) = prep_fields (init_env, raw_fields);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   515
    val envir_names = map fst envir;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   516
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   517
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   518
    (* args *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   519
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   520
    val defaultS = Sign.defaultS sign;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   521
    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
   522
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   523
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   524
    (* errors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   525
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   526
    val err_dup_parms =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   527
      (case duplicates params of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   528
        [] => []
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   529
      | dups => ["Duplicate parameters " ^ commas params]);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   530
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   531
    val err_extra_frees =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   532
      (case gen_rems (op =) (envir_names, params) of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   533
        [] => []
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   534
      | extras => ["Extraneous free type variables " ^ commas extras]);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   535
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   536
    val err_no_fields = if null bfields then ["No fields"] else [];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   537
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   538
    val err_dup_fields =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   539
      (case duplicates (map fst bfields) of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   540
        [] => []
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   541
      | dups => ["Duplicate fields " ^ commas_quote dups]);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   542
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   543
    val err_dup_sorts =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   544
      (case duplicates envir_names of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   545
        [] => []
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   546
      | dups => ["Inconsistent sort constraints for " ^ commas dups]);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   547
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   548
    val errs =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   549
      err_dup_parms @ err_extra_frees @ err_no_fields @ err_dup_fields @ err_dup_sorts;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   550
  in
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   551
    if null errs then ()
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   552
    else error (cat_lines errs);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   553
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   554
    writeln ("Defining record " ^ quote bname ^ " ...");
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   555
    thy |> record_definition (args, bname) parent parents bfields
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   556
  end
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   557
  handle ERROR => error ("Failed to define record " ^ quote bname);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   558
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   559
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
   560
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
   561
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   562
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   563
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   564
(** setup theory **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   565
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   566
val setup =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   567
 [Theory.init_data [record_thy_data],
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   568
  Theory.add_trfuns
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   569
    ([], [("_record", record_tr), ("_record_scheme", record_scheme_tr)], [], [])];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   570
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
end;