src/HOL/Tools/record_package.ML
author wenzelm
Fri, 30 Apr 1999 18:10:03 +0200
changeset 6556 daa00919502b
parent 6519 5bd1c469e742
child 6723 f342449d73ca
permissions -rw-r--r--
theory data: copy;
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.
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
     6
*)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
     7
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
     8
signature BASIC_RECORD_PACKAGE =
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
     9
sig
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    10
  val record_split_tac: int -> tactic
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
    11
  val record_split_name: string
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    12
  val record_split_wrapper: string * wrapper
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    13
end;
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
signature RECORD_PACKAGE =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    16
sig
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    17
  include BASIC_RECORD_PACKAGE
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    18
  val quiet_mode: bool ref
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    19
  val moreS: sort
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    20
  val mk_fieldT: (string * typ) * typ -> typ
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    21
  val dest_fieldT: typ -> (string * typ) * typ
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    22
  val mk_field: (string * term) * term -> term
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    23
  val mk_fst: term -> term
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    24
  val mk_snd: term -> term
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    25
  val mk_recordT: (string * typ) list * typ -> typ
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    26
  val dest_recordT: typ -> (string * typ) list * typ
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    27
  val mk_record: (string * term) list * term -> term
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    28
  val mk_sel: term -> string -> term
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    29
  val mk_update: term -> string * term -> term
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    30
  val print_records: theory -> unit
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    31
  val add_record: (string list * bstring) -> string option
6519
5bd1c469e742 iff_add_global (from simpdata.ML);
wenzelm
parents: 6394
diff changeset
    32
    -> (bstring * string) list -> theory -> theory * {simps: thm list, iffs: thm list}
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    33
  val add_record_i: (string list * bstring) -> (typ list * string) option
6519
5bd1c469e742 iff_add_global (from simpdata.ML);
wenzelm
parents: 6394
diff changeset
    34
    -> (bstring * typ) list -> theory -> theory * {simps: thm list, iffs: thm list}
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    35
  val setup: (theory -> theory) list
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    36
end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    37
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    38
structure RecordPackage: RECORD_PACKAGE =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    39
struct
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    40
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    41
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    42
(*** utilities ***)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    43
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    44
(* messages *)
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    45
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    46
val quiet_mode = ref false;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    47
fun message s = if ! quiet_mode then () else writeln s;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    48
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    49
6519
5bd1c469e742 iff_add_global (from simpdata.ML);
wenzelm
parents: 6394
diff changeset
    50
(* wrappers *)
5707
b0e631634b5a field_injects [iffs];
wenzelm
parents: 5698
diff changeset
    51
b0e631634b5a field_injects [iffs];
wenzelm
parents: 5698
diff changeset
    52
fun add_wrapper wrapper thy =
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
    53
  let val r = Classical.claset_ref_of thy
5707
b0e631634b5a field_injects [iffs];
wenzelm
parents: 5698
diff changeset
    54
  in r := ! r addSWrapper wrapper; thy end;
b0e631634b5a field_injects [iffs];
wenzelm
parents: 5698
diff changeset
    55
b0e631634b5a field_injects [iffs];
wenzelm
parents: 5698
diff changeset
    56
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    57
(* definitions and equations *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    58
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    59
infix 0 :== ===;
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    60
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    61
val (op :==) = Logic.mk_defpair;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    62
val (op ===) = HOLogic.mk_Trueprop o HOLogic.mk_eq;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    63
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
    64
fun get_defs thy specs = map (PureThy.get_thm thy o fst) specs;
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    65
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    66
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    67
(* proof by simplification *)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    68
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    69
fun prove_simp thy tacs simps =
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    70
  let
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
    71
    val sign = Theory.sign_of thy;
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
    72
    val ss = Simplifier.addsimps (HOL_basic_ss, simps);
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
    73
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
    74
    fun prove goal =
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
    75
      Goals.prove_goalw_cterm [] (Thm.cterm_of sign goal)
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
    76
        (K (tacs @ [ALLGOALS (Simplifier.simp_tac ss)]))
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
    77
      handle ERROR => error ("The error(s) above occurred while trying to prove "
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
    78
        ^ quote (Sign.string_of_term sign goal));
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    79
  in prove end;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    80
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    81
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    82
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    83
(*** syntax operations ***)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    84
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    85
(** name components **)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    86
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    87
val moreN = "more";
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    88
val schemeN = "_scheme";
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    89
val fieldN = "_field";
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    90
val raw_fieldN = "_raw_field";
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    91
val field_typeN = "_field_type";
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    92
val fstN = "_val";
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    93
val sndN = "_more";
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    94
val updateN = "_update";
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    95
val makeN = "make";
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    96
val make_schemeN = "make_scheme";
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    97
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    98
(*see datatype package*)
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    99
val caseN = "_case";
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   100
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   101
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   102
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   103
(** generic operations **)
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   104
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   105
(* adhoc priming of vars *)
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   106
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   107
fun prime (Free (x, T)) = Free (x ^ "'", T)
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   108
  | prime t = raise TERM ("prime: no free variable", [t]);
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   109
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   110
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   111
(* product case *)
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   112
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   113
fun fst_fn T U = Abs ("x", T, Abs ("y", U, Bound 1));
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   114
fun snd_fn T U = Abs ("x", T, Abs ("y", U, Bound 0));
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   115
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   116
fun mk_prod_case name f p =
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   117
  let
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   118
    val fT as Type ("fun", [A, Type ("fun", [B, C])]) = fastype_of f;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   119
    val pT = fastype_of p;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   120
  in Const (suffix caseN name, fT --> pT --> C) $ f $ p end;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   121
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   122
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   123
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   124
(** tuple operations **)
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
(* more type class *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   127
5210
54aaa779b6b4 removed global_names flag;
wenzelm
parents: 5201
diff changeset
   128
val moreS = ["Record.more"];
4867
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
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   131
(* types *)
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 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
   134
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   135
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
   136
      (case try (unsuffix field_typeN) c_field_type of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   137
        None => raise TYPE ("dest_fieldT", [typ], [])
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   138
      | Some c => ((c, T), U))
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   139
  | dest_fieldT typ = raise TYPE ("dest_fieldT", [typ], []);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   140
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   141
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   142
(* constructors *)
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
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
   145
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   146
fun gen_mk_field sfx ((c, t), u) =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   147
  let val T = fastype_of t and U = fastype_of u
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   148
  in Const (suffix sfx c, [T, U] ---> mk_fieldT ((c, T), U)) $ t $ u end;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   149
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   150
val mk_field = gen_mk_field fieldN;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   151
val mk_raw_field = gen_mk_field raw_fieldN;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   152
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
(* destructors *)
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 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
   157
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
   158
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   159
fun dest_field fst_or_snd p =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   160
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   161
    val pT = fastype_of p;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   162
    val ((c, T), U) = dest_fieldT pT;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   163
    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
   164
  in Const (suffix destN c, pT --> destT) $ p end;
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
val mk_fst = dest_field true;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   167
val mk_snd = dest_field false;
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
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   170
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   171
(** record operations **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   172
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   173
(* types *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   174
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   175
val mk_recordT = foldr mk_fieldT;
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 dest_recordT T =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   178
  (case try dest_fieldT T of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   179
    None => ([], T)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   180
  | 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
   181
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   182
fun find_fieldT c rT =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   183
  (case assoc (fst (dest_recordT rT), c) of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   184
    None => raise TYPE ("find_field: " ^ c, [rT], [])
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   185
  | Some T => T);
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
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   188
(* constructors *)
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
val mk_record = foldr mk_field;
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
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   193
(* selectors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   194
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   195
fun mk_selC rT (c, T) = (c, rT --> T);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   196
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   197
fun mk_sel r c =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   198
  let val rT = fastype_of r
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   199
  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
   200
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   201
val mk_moreC = mk_selC;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   202
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   203
fun mk_more r c =
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   204
  let val rT = fastype_of r
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   205
  in Const (mk_moreC rT (c, snd (dest_recordT rT))) $ r end;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   206
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   207
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   208
(* updates *)
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
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
   211
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   212
fun mk_update r (c, x) =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   213
  let val rT = fastype_of r
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   214
  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
   215
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   216
val mk_more_updateC = mk_updateC;
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   217
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   218
fun mk_more_update r (c, x) =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   219
  let val rT = fastype_of r
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   220
  in Const (mk_more_updateC rT (c, snd (dest_recordT rT))) $ x $ r end;
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   221
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   222
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   223
(* make *)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   224
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   225
fun mk_makeC rT (c, Ts) = (c, Ts ---> rT);
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   226
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   227
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   228
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   229
(** concrete syntax for records **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   230
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   231
(* parse translations *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   232
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   233
fun gen_field_tr mark sfx (t as Const (c, _) $ Free (name, _) $ arg) =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   234
      if c = mark then Syntax.const (suffix sfx name) $ arg
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   235
      else raise TERM ("gen_field_tr: " ^ mark, [t])
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   236
  | gen_field_tr mark _ t = raise TERM ("gen_field_tr: " ^ mark, [t]);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   237
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   238
fun gen_fields_tr sep mark sfx (tm as Const (c, _) $ t $ u) =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   239
      if c = sep then gen_field_tr mark sfx t :: gen_fields_tr sep mark sfx u
5201
wenzelm
parents: 5197
diff changeset
   240
      else [gen_field_tr mark sfx tm]
wenzelm
parents: 5197
diff changeset
   241
  | gen_fields_tr _ mark sfx tm = [gen_field_tr mark sfx tm];
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   242
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   243
fun gen_record_tr sep mark sfx unit [t] = foldr (op $) (gen_fields_tr sep mark sfx t, unit)
5201
wenzelm
parents: 5197
diff changeset
   244
  | gen_record_tr _ _ _ _ ts = raise TERM ("gen_record_tr", ts);
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   245
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   246
fun gen_record_scheme_tr sep mark sfx [t, more] = foldr (op $) (gen_fields_tr sep mark sfx t, more)
5201
wenzelm
parents: 5197
diff changeset
   247
  | gen_record_scheme_tr _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts);
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   248
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   249
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   250
val record_type_tr = gen_record_tr "_field_types" "_field_type" field_typeN (Syntax.const "unit");
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   251
val record_type_scheme_tr = gen_record_scheme_tr "_field_types" "_field_type" field_typeN;
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   252
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   253
val record_tr = gen_record_tr "_fields" "_field" fieldN HOLogic.unit;
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   254
val record_scheme_tr = gen_record_scheme_tr "_fields" "_field" fieldN;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   255
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   256
fun record_update_tr [t, u] =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   257
      foldr (op $) (rev (gen_fields_tr "_updates" "_update" updateN u), t)
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   258
  | record_update_tr ts = raise TERM ("record_update_tr", ts);
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   259
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   260
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   261
val parse_translation =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   262
 [("_record_type", record_type_tr),
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   263
  ("_record_type_scheme", record_type_scheme_tr),
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   264
  ("_record", record_tr),
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   265
  ("_record_scheme", record_scheme_tr),
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   266
  ("_record_update", record_update_tr)];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   267
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   268
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   269
(* print translations *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   270
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   271
fun gen_fields_tr' mark sfx (tm as Const (name_field, _) $ t $ u) =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   272
      (case try (unsuffix sfx) name_field of
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   273
        Some name =>
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   274
          apfst (cons (Syntax.const mark $ Syntax.free name $ t)) (gen_fields_tr' mark sfx u)
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   275
      | None => ([], tm))
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   276
  | gen_fields_tr' _ _ tm = ([], tm);
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   277
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   278
fun gen_record_tr' sep mark sfx is_unit record record_scheme tm =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   279
  let
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   280
    val (ts, u) = gen_fields_tr' mark sfx tm;
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   281
    val t' = foldr1 (fn (v, w) => Syntax.const sep $ v $ w) ts;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   282
  in
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   283
    if is_unit u then Syntax.const record $ t'
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   284
    else Syntax.const record_scheme $ t' $ u
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   285
  end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   286
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   287
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   288
val record_type_tr' =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   289
  gen_record_tr' "_field_types" "_field_type" field_typeN
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   290
    (fn Const ("unit", _) => true | _ => false) "_record_type" "_record_type_scheme";
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   291
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   292
val record_tr' =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   293
  gen_record_tr' "_fields" "_field" fieldN HOLogic.is_unit "_record" "_record_scheme";
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   294
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   295
fun record_update_tr' tm =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   296
  let val (ts, u) = gen_fields_tr' "_update" updateN tm in
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   297
    Syntax.const "_record_update" $ u $
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   298
      foldr1 (fn (v, w) => Syntax.const "_updates" $ v $ w) (rev ts)
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   299
  end;
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   300
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   301
5201
wenzelm
parents: 5197
diff changeset
   302
fun gen_field_tr' sfx tr' name =
wenzelm
parents: 5197
diff changeset
   303
  let val name_sfx = suffix sfx name
wenzelm
parents: 5197
diff changeset
   304
  in (name_sfx, fn [t, u] => tr' (Syntax.const name_sfx $ t $ u) | _ => raise Match) end;
wenzelm
parents: 5197
diff changeset
   305
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   306
fun print_translation names =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   307
  map (gen_field_tr' field_typeN record_type_tr') names @
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   308
  map (gen_field_tr' fieldN record_tr') names @
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   309
  map (gen_field_tr' updateN record_update_tr') names;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   310
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
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   313
(*** extend theory by record definition ***)
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
(** record info **)
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
(* type record_info and parent_info *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   318
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   319
type record_info =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   320
 {args: (string * sort) list,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   321
  parent: (typ list * string) option,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   322
  fields: (string * typ) list,
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
   323
  simps: thm list};
4867
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
type parent_info =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   326
 {name: string,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   327
  fields: (string * typ) list,
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
   328
  simps: thm list};
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
5052
bbe3584b515b fixed comment;
wenzelm
parents: 5006
diff changeset
   331
(* data kind 'HOL/records' *)
5001
9de7fda0a6df accomodate tuned version of theory data;
wenzelm
parents: 4970
diff changeset
   332
5006
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   333
structure RecordsArgs =
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   334
struct
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   335
  val name = "HOL/records";
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   336
  type T =
5707
b0e631634b5a field_injects [iffs];
wenzelm
parents: 5698
diff changeset
   337
    record_info Symtab.table *                          (*records*)
b0e631634b5a field_injects [iffs];
wenzelm
parents: 5698
diff changeset
   338
      (thm Symtab.table * Simplifier.simpset);          (*field split rules*)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   339
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   340
  val empty = (Symtab.empty, (Symtab.empty, HOL_basic_ss));
6556
daa00919502b theory data: copy;
wenzelm
parents: 6519
diff changeset
   341
  val copy = I;
5006
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   342
  val prep_ext = I;
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   343
  fun merge ((recs1, (sps1, ss1)), (recs2, (sps2, ss2))) =
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   344
    (Symtab.merge (K true) (recs1, recs2),
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   345
      (Symtab.merge (K true) (sps1, sps2), Simplifier.merge_ss (ss1, ss2)));
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   346
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   347
  fun print sg (recs, _) =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   348
    let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   349
      val prt_typ = Sign.pretty_typ sg;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   350
      val ext_const = Sign.cond_extern sg Sign.constK;
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
      fun pretty_parent None = []
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   353
        | pretty_parent (Some (Ts, name)) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   354
            [Pretty.block [prt_typ (Type (name, Ts)), Pretty.str " +"]];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   355
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   356
      fun pretty_field (c, T) = Pretty.block
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   357
        [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
   358
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   359
      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
   360
        (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
   361
          pretty_parent parent @ map pretty_field fields));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   362
    in
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   363
      seq (Pretty.writeln o pretty_record) (Symtab.dest recs)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   364
    end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   365
end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   366
5006
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   367
structure RecordsData = TheoryDataFun(RecordsArgs);
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   368
val print_records = RecordsData.print;
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   369
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   370
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   371
(* get and put data *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   372
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   373
fun get_record thy name = Symtab.lookup (#1 (RecordsData.get thy), name);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   374
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   375
fun put_record name info thy =
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   376
  let val (tab, sp) = RecordsData.get thy
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   377
  in RecordsData.put (Symtab.update ((name, info), tab), sp) thy end;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   378
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   379
fun add_record_splits splits thy =
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   380
  let
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   381
    val (tab, (sps, ss)) = RecordsData.get thy;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   382
    val simps = map #2 splits;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   383
  in RecordsData.put (tab, (Symtab.extend (sps, splits), Simplifier.addsimps (ss, simps))) thy end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   384
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   385
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   386
(* parent records *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   387
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   388
fun inst_record thy (types, name) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   389
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   390
    val sign = Theory.sign_of thy;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   391
    fun err msg = error (msg ^ " parent record " ^ quote name);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   392
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   393
    val {args, parent, fields, simps} =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   394
      (case get_record thy name of Some info => info | None => err "Unknown");
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   395
    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
   396
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   397
    fun bad_inst ((x, S), T) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   398
      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
   399
    val bads = mapfilter bad_inst (args ~~ types);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   400
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   401
    val inst = map fst args ~~ types;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   402
    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
   403
  in
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   404
    if not (null bads) then
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   405
      err ("Ill-sorted instantiation of " ^ commas bads ^ " in")
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   406
    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
   407
  end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   408
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   409
fun add_parents thy (None, parents) = parents
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   410
  | add_parents thy (Some (types, name), parents) =
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   411
      let val (pparent, pfields, psimps) = inst_record thy (types, name)
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   412
      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
   413
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
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   416
(** record field splitting **)
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   417
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   418
(* tactic *)
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   419
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   420
fun record_split_tac i st =
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   421
  let
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   422
    val (_, (sps, ss)) = RecordsData.get_sg (Thm.sign_of_thm st);
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   423
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   424
    fun is_fieldT (_, Type (a, [_, _])) = is_some (Symtab.lookup (sps, a))
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   425
      | is_fieldT _ = false;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   426
    val params = Logic.strip_params (Library.nth_elem (i - 1, Thm.prems_of st));
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   427
  in
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   428
    if exists is_fieldT params then Simplifier.full_simp_tac ss i st
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   429
    else Seq.empty
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   430
  end handle Library.LIST _ => Seq.empty;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   431
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   432
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   433
(* wrapper *)
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   434
5707
b0e631634b5a field_injects [iffs];
wenzelm
parents: 5698
diff changeset
   435
val record_split_name = "record_split_tac";
b0e631634b5a field_injects [iffs];
wenzelm
parents: 5698
diff changeset
   436
val record_split_wrapper = (record_split_name, fn tac => record_split_tac ORELSE' tac);
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   437
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   438
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   439
(* method *)
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   440
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   441
val record_split_method =
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   442
  ("record_split", Method.no_args (Method.METHOD0 (FIRSTGOAL record_split_tac)),
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   443
    "split record fields");
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   444
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   445
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   446
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   447
(** internal theory extenders **)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   448
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   449
(* field_type_defs *)
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   450
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   451
fun field_type_def ((thy, simps), (name, tname, vs, T, U)) =
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   452
  let
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6384
diff changeset
   453
    val full = Sign.full_name (Theory.sign_of thy);
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   454
    val (thy', {simps = simps', ...}) =
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   455
      thy
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   456
      |> setmp DatatypePackage.quiet_mode true
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   457
        (DatatypePackage.add_datatype_i true [tname]
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   458
          [(vs, tname, Syntax.NoSyn, [(name, [T, U], Syntax.NoSyn)])]);
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   459
    val thy'' =
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   460
      thy'
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   461
      |> setmp AxClass.quiet_mode true
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   462
        (AxClass.add_inst_arity_i (full tname, [HOLogic.termS, moreS], moreS) [] [] None);
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   463
  in (thy'', simps' @ simps) end;
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   464
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   465
fun field_type_defs args thy = foldl field_type_def ((thy, []), args);
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   466
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   467
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   468
(* field_definitions *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   469
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   470
fun field_definitions fields names zeta moreT more vars named_vars thy =
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   471
  let
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   472
    val sign = Theory.sign_of thy;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   473
    val base = Sign.base_name;
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   474
    val full_path = Sign.full_name_path sign;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   475
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   476
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   477
    (* prepare declarations and definitions *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   478
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   479
    (*field types*)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   480
    fun mk_fieldT_spec c =
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   481
      (suffix raw_fieldN c, suffix field_typeN c,
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   482
        ["'a", zeta], TFree ("'a", HOLogic.termS), moreT);
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   483
    val fieldT_specs = map (mk_fieldT_spec o base) names;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   484
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   485
    (*field constructors*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   486
    val field_decls = map (mk_fieldC moreT) fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   487
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   488
    fun mk_field_spec (c, v) =
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   489
      mk_field ((c, v), more) :== mk_raw_field ((c, v), more);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   490
    val field_specs = map mk_field_spec named_vars;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   491
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   492
    (*field destructors*)
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   493
    val dest_decls = map (mk_fstC moreT) fields @ map (mk_sndC moreT) fields;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   494
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   495
    fun mk_dest_spec dest f (c, T) =
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   496
      let val p = Free ("p", mk_fieldT ((c, T), moreT));
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   497
      in dest p :== mk_prod_case (suffix field_typeN c) (f T moreT) p end;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   498
    val dest_specs =
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   499
      map (mk_dest_spec mk_fst fst_fn) fields @
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   500
      map (mk_dest_spec mk_snd snd_fn) fields;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   501
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   502
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   503
    (* prepare theorems *)
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   504
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   505
    (*constructor injects*)
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   506
    fun mk_inject_prop (c, v) =
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   507
      HOLogic.mk_eq (mk_field ((c, v), more), mk_field ((c, prime v), prime more)) ===
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   508
        (HOLogic.conj $ HOLogic.mk_eq (v, prime v) $ HOLogic.mk_eq (more, prime more));
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   509
    val inject_props = map mk_inject_prop named_vars;
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   510
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   511
    (*destructor conversions*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   512
    fun mk_dest_prop dest dest' (c, v) =
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   513
      dest (mk_field ((c, v), more)) === dest' (v, more);
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   514
    val dest_props =
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   515
      map (mk_dest_prop mk_fst fst) named_vars @
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   516
      map (mk_dest_prop mk_snd snd) named_vars;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   517
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   518
    (*surjective pairing*)
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   519
    fun mk_surj_prop (c, T) =
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   520
      let val p = Free ("p", mk_fieldT ((c, T), moreT));
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   521
      in p === mk_field ((c, mk_fst p), mk_snd p) end;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   522
    val surj_props = map mk_surj_prop fields;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   523
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   524
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   525
    (* 1st stage: types_thy *)
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   526
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
   527
    val (types_thy, datatype_simps) =
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   528
      thy
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   529
      |> field_type_defs fieldT_specs;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   530
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   531
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   532
    (* 2nd stage: defs_thy *)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   533
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   534
    val defs_thy =
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   535
      types_thy
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   536
       |> (Theory.add_consts_i o map (Syntax.no_syn o apfst base))
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   537
         (field_decls @ dest_decls)
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
   538
       |> (PureThy.add_defs_i o map (fn x => (x, [Drule.tag_internal])))
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   539
         (field_specs @ dest_specs);
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   540
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   541
    val field_defs = get_defs defs_thy field_specs;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   542
    val dest_defs = get_defs defs_thy dest_specs;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   543
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   544
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   545
    (* 3rd stage: thms_thy *)
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   546
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   547
    val prove = prove_simp defs_thy;
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   548
    val prove_std = prove [] (field_defs @ dest_defs @ datatype_simps);
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   549
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   550
    val field_injects = map prove_std inject_props;
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   551
    val dest_convs = map prove_std dest_props;
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   552
    val surj_pairs = map (prove [DatatypePackage.induct_tac "p" 1]
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
   553
      (map Thm.symmetric field_defs @ dest_convs)) surj_props;
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   554
5707
b0e631634b5a field_injects [iffs];
wenzelm
parents: 5698
diff changeset
   555
    fun mk_split th = SplitPairedAll.rule (th RS eq_reflection);
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
   556
    val field_splits = map mk_split surj_pairs;
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   557
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   558
    val thms_thy =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   559
      defs_thy
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
   560
      |> (PureThy.add_thmss o map Thm.no_attributes)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   561
        [("field_defs", field_defs),
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   562
          ("dest_defs", dest_defs),
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   563
          ("dest_convs", dest_convs),
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   564
          ("surj_pairs", surj_pairs),
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   565
          ("field_splits", field_splits)];
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   566
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   567
  in (thms_thy, dest_convs, field_injects, field_splits) end;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   568
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   569
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   570
(* record_definition *)
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
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
   573
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   574
    val sign = Theory.sign_of thy;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   575
    val full = Sign.full_name_path sign bname;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   576
    val base = Sign.base_name;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   577
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   578
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   579
    (* basic components *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   580
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   581
    val alphas = map fst args;
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   582
    val name = Sign.full_name sign bname;       (*not made part of record name space!*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   583
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   584
    val parent_fields = flat (map #fields parents);
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   585
    val parent_names = map fst parent_fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   586
    val parent_types = map snd parent_fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   587
    val parent_len = length parent_fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   588
    val parent_xs = variantlist (map (base o fst) parent_fields, [moreN]);
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   589
    val parent_vars = ListPair.map Free (parent_xs, parent_types);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   590
    val parent_named_vars = parent_names ~~ parent_vars;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   591
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   592
    val fields = map (apfst full) bfields;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   593
    val names = map fst fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   594
    val types = map snd fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   595
    val len = length fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   596
    val xs = variantlist (map fst bfields, moreN :: parent_xs);
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   597
    val vars = ListPair.map Free (xs, types);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   598
    val named_vars = names ~~ vars;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   599
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   600
    val all_fields = parent_fields @ fields;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   601
    val all_names = parent_names @ names;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   602
    val all_types = parent_types @ types;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   603
    val all_len = parent_len + len;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   604
    val all_xs = parent_xs @ xs;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   605
    val all_vars = parent_vars @ vars;
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   606
    val all_named_vars = parent_named_vars @ named_vars;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   607
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   608
    val zeta = variant alphas "'z";
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   609
    val moreT = TFree (zeta, moreS);
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   610
    val more = Free (moreN, moreT);
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   611
    val full_moreN = full moreN;
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   612
    fun more_part t = mk_more t full_moreN;
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   613
    fun more_part_update t x = mk_more_update t (full_moreN, x);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   614
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   615
    val parent_more = funpow parent_len mk_snd;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   616
    val idxs = 0 upto (len - 1);
4867
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
    val rec_schemeT = mk_recordT (all_fields, moreT);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   619
    val rec_scheme = mk_record (all_named_vars, more);
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   620
    val r = Free ("r", rec_schemeT);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   621
    val recT = mk_recordT (all_fields, HOLogic.unitT);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   622
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   623
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   624
    (* prepare print translation functions *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   625
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   626
    val field_tr's =
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   627
      print_translation (distinct (flat (map NameSpace.accesses (full_moreN :: names))));
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   628
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   629
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   630
    (* prepare declarations *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   631
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   632
    val sel_decls = map (mk_selC rec_schemeT) bfields @
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   633
      [mk_moreC rec_schemeT (moreN, moreT)];
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   634
    val update_decls = map (mk_updateC rec_schemeT) bfields @
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   635
      [mk_more_updateC rec_schemeT (moreN, moreT)];
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   636
    val make_decls =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   637
      [(mk_makeC rec_schemeT (make_schemeN, all_types @ [moreT])),
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   638
       (mk_makeC recT (makeN, all_types))];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   639
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   640
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   641
    (* prepare definitions *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   642
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   643
    (*record (scheme) type abbreviation*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   644
    val recordT_specs =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   645
      [(suffix schemeN bname, alphas @ [zeta], rec_schemeT, Syntax.NoSyn),
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   646
        (bname, alphas, recT, Syntax.NoSyn)];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   647
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   648
    (*selectors*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   649
    fun mk_sel_spec (i, c) =
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   650
      mk_sel r c :== mk_fst (funpow i mk_snd (parent_more r));
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   651
    val sel_specs =
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   652
      ListPair.map mk_sel_spec (idxs, names) @
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   653
        [more_part r :== funpow len mk_snd (parent_more r)];
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   654
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   655
    (*updates*)
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   656
    val all_sels = all_names ~~ map (mk_sel r) all_names;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   657
    fun mk_upd_spec (i, (c, x)) =
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   658
      mk_update r (c, x) :==
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   659
        mk_record (nth_update (c, x) (parent_len + i, all_sels), more_part r)
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   660
    val update_specs =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   661
      ListPair.map mk_upd_spec (idxs, named_vars) @
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   662
        [more_part_update r more :== mk_record (all_sels, more)];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   663
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   664
    (*makes*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   665
    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
   666
    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
   667
    val make_specs =
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   668
      [list_comb (make_scheme, all_vars) $ more :== rec_scheme,
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   669
        list_comb (make, all_vars) :== mk_record (all_named_vars, HOLogic.unit)];
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   670
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   671
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   672
    (* prepare propositions *)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   673
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   674
    (*selectors*)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   675
    val sel_props =
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   676
      map (fn (c, x) => mk_sel rec_scheme c === x) named_vars @
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   677
        [more_part rec_scheme === more];
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   678
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   679
    (*updates*)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   680
    fun mk_upd_prop (i, (c, T)) =
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   681
      let val x' = Free (variant all_xs (base c ^ "'"), T) in
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   682
        mk_update rec_scheme (c, x') ===
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   683
          mk_record (nth_update (c, x') (parent_len + i, all_named_vars), more)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   684
      end;
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   685
    val update_props =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   686
      ListPair.map mk_upd_prop (idxs, fields) @
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   687
        let val more' = Free (variant all_xs (moreN ^ "'"), moreT)
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   688
        in [more_part_update rec_scheme more' === mk_record (all_named_vars, more')] end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   689
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   690
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   691
    (* 1st stage: fields_thy *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   692
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   693
    val (fields_thy, field_simps, field_injects, field_splits) =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   694
      thy
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   695
      |> Theory.add_path bname
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   696
      |> 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
   697
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
   698
    val named_splits = map2 (fn (c, th) => (suffix field_typeN c, th)) (names, field_splits);
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   699
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   700
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   701
    (* 2nd stage: defs_thy *)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   702
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   703
    val defs_thy =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   704
      fields_thy
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   705
      |> Theory.parent_path
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   706
      |> Theory.add_tyabbrs_i recordT_specs     (*not made part of record name space!*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   707
      |> Theory.add_path bname
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   708
      |> Theory.add_trfuns ([], [], field_tr's, [])
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   709
      |> (Theory.add_consts_i o map Syntax.no_syn)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   710
        (sel_decls @ update_decls @ make_decls)
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
   711
      |> (PureThy.add_defs_i o map (fn x => (x, [Drule.tag_internal])))
5212
2bc5b5cf0516 make_defs not marked as internal;
wenzelm
parents: 5210
diff changeset
   712
        (sel_specs @ update_specs)
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
   713
      |> (PureThy.add_defs_i o map Thm.no_attributes) make_specs;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   714
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   715
    val sel_defs = get_defs defs_thy sel_specs;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   716
    val update_defs = get_defs defs_thy update_specs;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   717
    val make_defs = get_defs defs_thy make_specs;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   718
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   719
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   720
    (* 3rd stage: thms_thy *)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   721
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   722
    val parent_simps = flat (map #simps parents);
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   723
    val prove = prove_simp defs_thy [];
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   724
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   725
    val sel_convs = map (prove (parent_simps @ sel_defs @ field_simps)) sel_props;
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   726
    val update_convs = map (prove (parent_simps @ update_defs @ sel_convs)) update_props;
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   727
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   728
    val simps = field_simps @ sel_convs @ update_convs @ make_defs;
6519
5bd1c469e742 iff_add_global (from simpdata.ML);
wenzelm
parents: 6394
diff changeset
   729
    val iffs = field_injects;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   730
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   731
    val thms_thy =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   732
      defs_thy
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
   733
      |> (PureThy.add_thmss o map Thm.no_attributes)
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   734
        [("select_defs", sel_defs),
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   735
          ("update_defs", update_defs),
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   736
          ("make_defs", make_defs),
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   737
          ("select_convs", sel_convs),
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   738
          ("update_convs", update_convs)]
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
   739
      |> PureThy.add_thmss
5707
b0e631634b5a field_injects [iffs];
wenzelm
parents: 5698
diff changeset
   740
        [(("simps", simps), [Simplifier.simp_add_global]),
6519
5bd1c469e742 iff_add_global (from simpdata.ML);
wenzelm
parents: 6394
diff changeset
   741
         (("iffs", iffs), [iff_add_global])];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   742
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   743
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   744
    (* 4th stage: final_thy *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   745
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   746
    val final_thy =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   747
      thms_thy
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   748
      |> put_record name {args = args, parent = parent, fields = fields, simps = simps}
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   749
      |> add_record_splits named_splits
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   750
      |> Theory.parent_path;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   751
6519
5bd1c469e742 iff_add_global (from simpdata.ML);
wenzelm
parents: 6394
diff changeset
   752
  in (final_thy, {simps = simps, iffs = iffs}) end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   753
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   754
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   755
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   756
(** theory extender interface **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   757
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   758
(* prepare arguments *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   759
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   760
(*note: read_raw_typ avoids expanding type abbreviations*)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   761
fun read_raw_parent sign s =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   762
  (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
   763
    Type (name, Ts) => (Ts, name)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   764
  | _ => error ("Bad parent record specification: " ^ quote s));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   765
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   766
fun read_typ sign (env, s) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   767
  let
5060
7b86df67cc1a def_sort;
wenzelm
parents: 5052
diff changeset
   768
    fun def_sort (x, ~1) = assoc (env, x)
7b86df67cc1a def_sort;
wenzelm
parents: 5052
diff changeset
   769
      | def_sort _ = None;
7b86df67cc1a def_sort;
wenzelm
parents: 5052
diff changeset
   770
    val T = Type.no_tvars (Sign.read_typ (sign, def_sort) s) handle TYPE (msg, _, _) => error msg;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   771
  in (Term.add_typ_tfrees (T, env), T) end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   772
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   773
fun cert_typ sign (env, raw_T) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   774
  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
   775
  in (Term.add_typ_tfrees (T, env), T) end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   776
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   777
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   778
(* add_record *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   779
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   780
(*we do all preparations and error checks here, deferring the real
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   781
  work to record_definition*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   782
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   783
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
   784
  let
4970
8b65444edbb0 Changed require to requires for MLWorks
paulson
parents: 4967
diff changeset
   785
    val _ = Theory.requires thy "Record" "record definitions";
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   786
    val sign = Theory.sign_of thy;
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   787
    val _ = message ("Defining record " ^ quote bname ^ " ...");
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   788
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   789
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   790
    (* parents *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   791
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   792
    fun prep_inst T = snd (cert_typ sign ([], T));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   793
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   794
    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
   795
      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
   796
    val parents = add_parents thy (parent, []);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   797
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   798
    val init_env =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   799
      (case parent of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   800
        None => []
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   801
      | Some (types, _) => foldr Term.add_typ_tfrees (types, []));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   802
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   803
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   804
    (* fields *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   805
4967
c9c481bc0216 foldl_map prep_field;
wenzelm
parents: 4895
diff changeset
   806
    fun prep_field (env, (c, raw_T)) =
c9c481bc0216 foldl_map prep_field;
wenzelm
parents: 4895
diff changeset
   807
      let val (env', T) = prep_typ sign (env, raw_T) handle ERROR =>
c9c481bc0216 foldl_map prep_field;
wenzelm
parents: 4895
diff changeset
   808
        error ("The error(s) above occured in field " ^ quote c)
c9c481bc0216 foldl_map prep_field;
wenzelm
parents: 4895
diff changeset
   809
      in (env', (c, T)) end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   810
4967
c9c481bc0216 foldl_map prep_field;
wenzelm
parents: 4895
diff changeset
   811
    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
   812
    val envir_names = map fst envir;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   813
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   814
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   815
    (* args *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   816
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   817
    val defaultS = Sign.defaultS sign;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   818
    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
   819
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   820
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   821
    (* errors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   822
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   823
    val name = Sign.full_name sign bname;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   824
    val err_dup_record =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   825
      if is_none (get_record thy name) then []
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   826
      else ["Duplicate definition of record " ^ quote name];
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   827
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   828
    val err_dup_parms =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   829
      (case duplicates params of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   830
        [] => []
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   831
      | dups => ["Duplicate parameter(s) " ^ commas dups]);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   832
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   833
    val err_extra_frees =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   834
      (case gen_rems (op =) (envir_names, params) of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   835
        [] => []
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   836
      | extras => ["Extra free type variable(s) " ^ commas extras]);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   837
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   838
    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
   839
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   840
    val err_dup_fields =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   841
      (case duplicates (map fst bfields) of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   842
        [] => []
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   843
      | dups => ["Duplicate field(s) " ^ commas_quote dups]);
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   844
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   845
    val err_bad_fields =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   846
      if forall (not_equal moreN o fst) bfields then []
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   847
      else ["Illegal field name " ^ quote moreN];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   848
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   849
    val err_dup_sorts =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   850
      (case duplicates envir_names of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   851
        [] => []
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   852
      | dups => ["Inconsistent sort constraints for " ^ commas dups]);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   853
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   854
    val errs =
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   855
      err_dup_record @ err_dup_parms @ err_extra_frees @ err_no_fields @
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   856
      err_dup_fields @ err_bad_fields @ err_dup_sorts;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   857
  in
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   858
    if null errs then () else error (cat_lines errs);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   859
    thy |> record_definition (args, bname) parent parents bfields
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   860
  end
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   861
  handle ERROR => error ("Failed to define record " ^ quote bname);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   862
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   863
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
   864
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
   865
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   866
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   867
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   868
(** package setup **)
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   869
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   870
(* setup theory *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   871
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   872
val setup =
5006
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   873
 [RecordsData.init,
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   874
  Theory.add_trfuns ([], parse_translation, [], []),
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   875
  Method.add_methods [record_split_method],
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   876
  add_wrapper record_split_wrapper];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   877
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   878
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   879
(* outer syntax *)
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   880
6384
eed1273c9146 local open OuterParse;
wenzelm
parents: 6358
diff changeset
   881
local open OuterParse in
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   882
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   883
val record_decl =
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   884
  type_args -- name -- ($$$ "=" |-- Scan.option (typ --| $$$ "+")
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   885
    -- Scan.repeat1 (name -- ($$$ "::" |-- typ)));
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   886
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   887
val recordP =
6384
eed1273c9146 local open OuterParse;
wenzelm
parents: 6358
diff changeset
   888
  OuterSyntax.command "record" "define extensible record"
6519
5bd1c469e742 iff_add_global (from simpdata.ML);
wenzelm
parents: 6394
diff changeset
   889
    (record_decl >> (fn (x, (y, z)) => Toplevel.theory (#1 o add_record x y z)));
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   890
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   891
val _ = OuterSyntax.add_parsers [recordP];
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   892
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   893
end;
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   894
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   895
6384
eed1273c9146 local open OuterParse;
wenzelm
parents: 6358
diff changeset
   896
end;
eed1273c9146 local open OuterParse;
wenzelm
parents: 6358
diff changeset
   897
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   898
structure BasicRecordPackage: BASIC_RECORD_PACKAGE = RecordPackage;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   899
open BasicRecordPackage;