src/HOL/Tools/record_package.ML
author wenzelm
Sat, 27 Oct 2001 00:06:22 +0200
changeset 11959 ed914e8a0fd1
parent 11940 80365073b8b3
child 11967 49c7f03cd311
permissions -rw-r--r--
exclude field_simps from user-level "simps";
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
9230
wenzelm
parents: 9040
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
     5
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
     6
Extensible records with structural subtyping in HOL.
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
     7
*)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
     8
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
     9
signature BASIC_RECORD_PACKAGE =
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    10
sig
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
    11
  val record_simproc: simproc
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    12
  val record_split_tac: int -> tactic
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
    13
  val record_split_name: string
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    14
  val record_split_wrapper: string * wrapper
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    15
end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    16
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    17
signature RECORD_PACKAGE =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    18
sig
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    19
  include BASIC_RECORD_PACKAGE
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    20
  val quiet_mode: bool ref
8574
bed3b994ab26 export updateN;
wenzelm
parents: 8428
diff changeset
    21
  val updateN: string
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    22
  val mk_fieldT: (string * typ) * typ -> typ
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    23
  val dest_fieldT: typ -> (string * typ) * typ
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    24
  val mk_field: (string * term) * term -> term
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    25
  val mk_fst: term -> term
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    26
  val mk_snd: term -> term
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    27
  val mk_recordT: (string * typ) list * typ -> typ
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    28
  val dest_recordT: typ -> (string * typ) list * typ
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    29
  val mk_record: (string * term) list * term -> term
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    30
  val mk_sel: term -> string -> term
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    31
  val mk_update: term -> string * term -> term
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    32
  val print_records: theory -> unit
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    33
  val add_record: (string list * bstring) -> string option
6519
5bd1c469e742 iff_add_global (from simpdata.ML);
wenzelm
parents: 6394
diff changeset
    34
    -> (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
    35
  val add_record_i: (string list * bstring) -> (typ list * string) option
6519
5bd1c469e742 iff_add_global (from simpdata.ML);
wenzelm
parents: 6394
diff changeset
    36
    -> (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
    37
  val setup: (theory -> theory) list
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    38
end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    39
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    40
structure RecordPackage: RECORD_PACKAGE =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    41
struct
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    42
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    43
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
    44
(*** theory context references ***)
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
    45
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
    46
val product_typeN = "Record.product_type";
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
    47
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
    48
val product_typeI = thm "product_typeI";
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
    49
val product_type_inject = thm "product_type_inject";
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
    50
val product_type_conv1 = thm "product_type_conv1";
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
    51
val product_type_conv2 = thm "product_type_conv2";
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    52
val product_type_induct = thm "product_type_induct";
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    53
val product_type_cases = thm "product_type_cases";
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
    54
val product_type_split_paired_all = thm "product_type_split_paired_all";
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
    55
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
    56
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
    57
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    58
(*** utilities ***)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    59
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    60
(* messages *)
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    61
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    62
val quiet_mode = ref false;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    63
fun message s = if ! quiet_mode then () else writeln s;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    64
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    65
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
    66
(* fundamental syntax *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    67
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
    68
fun prefix_base s = NameSpace.map_base (fn bname => s ^ bname);
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
    69
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    70
val Trueprop = HOLogic.mk_Trueprop;
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    71
fun All xs t = Term.list_all_free (xs, t);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    72
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
    73
infix 9 $$;
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
    74
infix 0 :== ===;
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
    75
infixr 0 ==>;
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
    76
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
    77
val (op $$) = Term.list_comb;
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    78
val (op :==) = Logic.mk_defpair;
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    79
val (op ===) = Trueprop o HOLogic.mk_eq;
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    80
val (op ==>) = Logic.mk_implies;
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    81
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    82
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    83
(* proof tools *)
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    84
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    85
fun prove_goal sign goal tacs =
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    86
  Goals.prove_goalw_cterm [] (Thm.cterm_of sign goal) tacs
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    87
  handle ERROR => error ("The error(s) above occurred while trying to prove " ^
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    88
    quote (Sign.string_of_term sign goal));
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
    89
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
    90
fun prove_simp sign ss tacs simps =
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    91
  let
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
    92
    val ss' = Simplifier.addsimps (ss, simps);
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    93
    fun prove goal = prove_goal sign goal
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    94
      (K (tacs @ [ALLGOALS (Simplifier.asm_full_simp_tac ss')]));
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    95
  in prove end;
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
    96
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    97
fun try_param_tac x s rule i st =
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
    98
  res_inst_tac [(x, (case Tactic.innermost_params i st of [] => s | (p, _) :: _ => p))] rule i st;
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
    99
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   100
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   101
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   102
(*** syntax operations ***)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   103
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   104
(** name components **)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   105
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   106
val rN = "r";
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   107
val moreN = "more";
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   108
val schemeN = "_scheme";
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   109
val field_typeN = "_field_type";
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   110
val fieldN = "_field";
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   111
val fstN = "_val";
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   112
val sndN = "_more";
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   113
val updateN = "_update";
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   114
val makeN = "make";
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   115
val extendN = "extend";
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   116
val truncateN = "truncate";
11940
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   117
val fieldsN = "fields";
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   118
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   119
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   120
(*see typedef_package.ML*)
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   121
val RepN = "Rep_";
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   122
val AbsN = "Abs_";
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   123
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   124
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   125
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   126
(** tuple operations **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   127
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   128
(* types *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   129
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   130
fun mk_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
   131
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   132
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
   133
      (case try (unsuffix field_typeN) c_field_type of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   134
        None => raise TYPE ("dest_fieldT", [typ], [])
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   135
      | Some c => ((c, T), U))
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   136
  | dest_fieldT typ = raise TYPE ("dest_fieldT", [typ], []);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   137
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   138
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   139
(* morphisms *)
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   140
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   141
fun mk_Rep U (c, T) =
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   142
  Const (suffix field_typeN (prefix_base RepN c),
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   143
    mk_fieldT ((c, T), U) --> HOLogic.mk_prodT (T, U));
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   144
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   145
fun mk_Abs U (c, T) =
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   146
  Const (suffix field_typeN (prefix_base AbsN c),
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   147
    HOLogic.mk_prodT (T, U) --> mk_fieldT ((c, T), U));
11833
wenzelm
parents: 11832
diff changeset
   148
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   149
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   150
(* constructors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   151
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   152
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
   153
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   154
fun mk_field ((c, t), u) =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   155
  let val T = fastype_of t and U = fastype_of u
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   156
  in Const (suffix fieldN c, [T, U] ---> mk_fieldT ((c, T), U)) $ t $ u end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   157
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   158
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   159
(* destructors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   160
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   161
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
   162
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
   163
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   164
fun dest_field fst_or_snd p =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   165
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   166
    val pT = fastype_of p;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   167
    val ((c, T), U) = dest_fieldT pT;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   168
    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
   169
  in Const (suffix destN c, pT --> destT) $ p end;
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
val mk_fst = dest_field true;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   172
val mk_snd = dest_field false;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   173
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
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   176
(** record operations **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   177
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   178
(* types *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   179
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   180
val mk_recordT = foldr mk_fieldT;
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 dest_recordT T =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   183
  (case try dest_fieldT T of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   184
    None => ([], T)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   185
  | 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
   186
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   187
fun find_fieldT c rT =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   188
  (case assoc (fst (dest_recordT rT), c) of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   189
    None => raise TYPE ("find_field: " ^ c, [rT], [])
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   190
  | Some T => T);
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
(* constructors *)
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
val mk_record = foldr mk_field;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   196
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   197
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   198
(* selectors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   199
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   200
fun mk_selC rT (c, T) = (c, rT --> T);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   201
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   202
fun mk_sel r c =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   203
  let val rT = fastype_of r
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   204
  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
   205
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   206
fun mk_named_sels names r = names ~~ map (mk_sel r) names;
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   207
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   208
val mk_moreC = mk_selC;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   209
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   210
fun mk_more r c =
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   211
  let val rT = fastype_of r
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   212
  in Const (mk_moreC rT (c, snd (dest_recordT rT))) $ r end;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   213
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   214
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   215
(* updates *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   216
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   217
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
   218
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   219
fun mk_update r (c, x) =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   220
  let val rT = fastype_of r
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   221
  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
   222
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   223
val mk_more_updateC = mk_updateC;
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   224
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   225
fun mk_more_update r (c, x) =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   226
  let val rT = fastype_of r
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   227
  in Const (mk_more_updateC rT (c, snd (dest_recordT rT))) $ x $ r end;
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   228
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   229
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   230
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   231
(** concrete syntax for records **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   232
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   233
(* parse translations *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   234
11473
4546d8d39221 fix problem with user translations by making field names appear as consts;
wenzelm
parents: 10008
diff changeset
   235
fun gen_field_tr mark sfx (t as Const (c, _) $ Const (name, _) $ arg) =
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   236
      if c = mark then Syntax.const (suffix sfx name) $ arg
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   237
      else raise TERM ("gen_field_tr: " ^ mark, [t])
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   238
  | 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
   239
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   240
fun gen_fields_tr sep mark sfx (tm as Const (c, _) $ t $ u) =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   241
      if c = sep then gen_field_tr mark sfx t :: gen_fields_tr sep mark sfx u
5201
wenzelm
parents: 5197
diff changeset
   242
      else [gen_field_tr mark sfx tm]
wenzelm
parents: 5197
diff changeset
   243
  | gen_fields_tr _ mark sfx tm = [gen_field_tr mark sfx tm];
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   244
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   245
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
   246
  | gen_record_tr _ _ _ _ ts = raise TERM ("gen_record_tr", ts);
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   247
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   248
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
   249
  | gen_record_scheme_tr _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts);
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   250
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   251
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   252
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
   253
val record_type_scheme_tr = gen_record_scheme_tr "_field_types" "_field_type" field_typeN;
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   254
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   255
val record_tr = gen_record_tr "_fields" "_field" fieldN HOLogic.unit;
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   256
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
   257
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   258
fun record_update_tr [t, u] =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   259
      foldr (op $) (rev (gen_fields_tr "_updates" "_update" updateN u), t)
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   260
  | record_update_tr ts = raise TERM ("record_update_tr", ts);
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   261
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   262
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   263
fun update_name_tr (Free (x, T) :: ts) = Free (suffix updateN x, T) $$ ts
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   264
  | update_name_tr (Const (x, T) :: ts) = Const (suffix updateN x, T) $$ ts
11833
wenzelm
parents: 11832
diff changeset
   265
  | update_name_tr (((c as Const ("_constrain", _)) $ t $ ty) :: ts) =
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   266
      (c $ update_name_tr [t] $ (Syntax.const "fun" $ ty $ Syntax.const "dummy")) $$ ts
11833
wenzelm
parents: 11832
diff changeset
   267
  | update_name_tr ts = raise TERM ("update_name_tr", ts);
wenzelm
parents: 11832
diff changeset
   268
wenzelm
parents: 11832
diff changeset
   269
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   270
val parse_translation =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   271
 [("_record_type", record_type_tr),
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   272
  ("_record_type_scheme", record_type_scheme_tr),
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   273
  ("_record", record_tr),
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   274
  ("_record_scheme", record_scheme_tr),
11833
wenzelm
parents: 11832
diff changeset
   275
  ("_record_update", record_update_tr),
wenzelm
parents: 11832
diff changeset
   276
  ("_update_name", update_name_tr)];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   277
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   278
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   279
(* print translations *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   280
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   281
fun gen_fields_tr' mark sfx (tm as Const (name_field, _) $ t $ u) =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   282
      (case try (unsuffix sfx) name_field of
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   283
        Some name =>
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   284
          apfst (cons (Syntax.const mark $ Syntax.free name $ t)) (gen_fields_tr' mark sfx u)
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   285
      | None => ([], tm))
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   286
  | gen_fields_tr' _ _ tm = ([], tm);
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   287
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   288
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
   289
  let
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   290
    val (ts, u) = gen_fields_tr' mark sfx tm;
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   291
    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
   292
  in
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   293
    if is_unit u then Syntax.const record $ t'
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   294
    else Syntax.const record_scheme $ t' $ u
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   295
  end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   296
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   297
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   298
val record_type_tr' =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   299
  gen_record_tr' "_field_types" "_field_type" field_typeN
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   300
    (fn Const ("unit", _) => true | _ => false) "_record_type" "_record_type_scheme";
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   301
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   302
val record_tr' =
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   303
  gen_record_tr' "_fields" "_field" fieldN
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   304
    (fn Const ("Unity", _) => true | _ => false) "_record" "_record_scheme";
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   305
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   306
fun record_update_tr' tm =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   307
  let val (ts, u) = gen_fields_tr' "_update" updateN tm in
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   308
    Syntax.const "_record_update" $ u $
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   309
      foldr1 (fn (v, w) => Syntax.const "_updates" $ v $ w) (rev ts)
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   310
  end;
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   311
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   312
5201
wenzelm
parents: 5197
diff changeset
   313
fun gen_field_tr' sfx tr' name =
wenzelm
parents: 5197
diff changeset
   314
  let val name_sfx = suffix sfx name
wenzelm
parents: 5197
diff changeset
   315
  in (name_sfx, fn [t, u] => tr' (Syntax.const name_sfx $ t $ u) | _ => raise Match) end;
wenzelm
parents: 5197
diff changeset
   316
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   317
fun print_translation names =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   318
  map (gen_field_tr' field_typeN record_type_tr') names @
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   319
  map (gen_field_tr' fieldN record_tr') names @
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   320
  map (gen_field_tr' updateN record_update_tr') names;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   321
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   322
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   323
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   324
(*** extend theory by record definition ***)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   325
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   326
(** record info **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   327
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   328
(* type record_info and parent_info *)
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
type record_info =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   331
 {args: (string * sort) list,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   332
  parent: (typ list * string) option,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   333
  fields: (string * typ) list,
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   334
  simps: thm list, induct: thm, cases: thm};
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   335
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   336
fun make_record_info args parent fields simps induct cases =
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   337
 {args = args, parent = parent, fields = fields, simps = simps,
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   338
  induct = induct, cases = cases}: record_info;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   339
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   340
type parent_info =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   341
 {name: string,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   342
  fields: (string * typ) list,
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   343
  simps: thm list, induct: thm, cases: thm};
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   344
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   345
fun make_parent_info name fields simps induct cases =
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   346
 {name = name, fields = fields, simps = simps,
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   347
  induct = induct, cases = cases}: parent_info;
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   348
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   349
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   350
5052
bbe3584b515b fixed comment;
wenzelm
parents: 5006
diff changeset
   351
(* data kind 'HOL/records' *)
5001
9de7fda0a6df accomodate tuned version of theory data;
wenzelm
parents: 4970
diff changeset
   352
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   353
type record_data =
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   354
 {records: record_info Symtab.table,
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   355
  sel_upd:
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   356
   {selectors: unit Symtab.table,
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   357
    updates: string Symtab.table,
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   358
    simpset: Simplifier.simpset},
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   359
  field_splits:
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   360
   {fields: unit Symtab.table,
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   361
    simpset: Simplifier.simpset}};
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   362
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   363
fun make_record_data records sel_upd field_splits =
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   364
 {records = records, sel_upd = sel_upd, field_splits = field_splits}: record_data;
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   365
5006
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   366
structure RecordsArgs =
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   367
struct
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   368
  val name = "HOL/records";
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   369
  type T = record_data;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   370
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   371
  val empty =
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   372
    make_record_data Symtab.empty
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   373
      {selectors = Symtab.empty, updates = Symtab.empty, simpset = HOL_basic_ss}
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   374
      {fields = Symtab.empty, simpset = HOL_basic_ss};
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   375
6556
daa00919502b theory data: copy;
wenzelm
parents: 6519
diff changeset
   376
  val copy = I;
5006
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   377
  val prep_ext = I;
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   378
  fun merge
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   379
   ({records = recs1,
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   380
     sel_upd = {selectors = sels1, updates = upds1, simpset = ss1},
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   381
     field_splits = {fields = flds1, simpset = fld_ss1}},
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   382
    {records = recs2,
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   383
     sel_upd = {selectors = sels2, updates = upds2, simpset = ss2},
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   384
     field_splits = {fields = flds2, simpset = fld_ss2}}) =
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   385
    make_record_data
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   386
      (Symtab.merge (K true) (recs1, recs2))
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   387
      {selectors = Symtab.merge (K true) (sels1, sels2),
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   388
        updates = Symtab.merge (K true) (upds1, upds2),
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   389
        simpset = Simplifier.merge_ss (ss1, ss2)}
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   390
      {fields = Symtab.merge (K true) (flds1, flds2),
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   391
        simpset = Simplifier.merge_ss (fld_ss1, fld_ss2)};
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   392
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   393
  fun print sg ({records = recs, ...}: record_data) =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   394
    let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   395
      val prt_typ = Sign.pretty_typ sg;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   396
      val ext_const = Sign.cond_extern sg Sign.constK;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   397
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   398
      fun pretty_parent None = []
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   399
        | pretty_parent (Some (Ts, name)) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   400
            [Pretty.block [prt_typ (Type (name, Ts)), Pretty.str " +"]];
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   401
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   402
      fun pretty_field (c, T) = Pretty.block
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   403
        [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
   404
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   405
      fun pretty_record (name, {args, parent, fields, simps = _, induct = _, cases = _}) =
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   406
        Pretty.block (Pretty.fbreaks (Pretty.block
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   407
          [prt_typ (Type (name, map TFree args)), Pretty.str " = "] ::
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   408
          pretty_parent parent @ map pretty_field fields));
8720
840c75ab2a7f Pretty.chunks;
wenzelm
parents: 8574
diff changeset
   409
    in
840c75ab2a7f Pretty.chunks;
wenzelm
parents: 8574
diff changeset
   410
      map pretty_record (Sign.cond_extern_table sg Sign.typeK recs)
840c75ab2a7f Pretty.chunks;
wenzelm
parents: 8574
diff changeset
   411
      |> Pretty.chunks |> Pretty.writeln
840c75ab2a7f Pretty.chunks;
wenzelm
parents: 8574
diff changeset
   412
    end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   413
end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   414
5006
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   415
structure RecordsData = TheoryDataFun(RecordsArgs);
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   416
val print_records = RecordsData.print;
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
   417
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   418
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   419
(* access 'records' *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   420
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   421
fun get_record thy name = Symtab.lookup (#records (RecordsData.get thy), name);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   422
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   423
fun put_record name info thy =
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   424
  let
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   425
    val {records, sel_upd, field_splits} = RecordsData.get thy;
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   426
    val data = make_record_data (Symtab.update ((name, info), records)) sel_upd field_splits;
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   427
  in RecordsData.put data thy end;
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   428
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   429
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   430
(* access 'sel_upd' *)
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   431
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   432
fun get_sel_upd sg = #sel_upd (RecordsData.get_sg sg);
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   433
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   434
fun get_selectors sg name = Symtab.lookup (#selectors (get_sel_upd sg), name);
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   435
fun get_updates sg name = Symtab.lookup (#updates (get_sel_upd sg), name);
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   436
fun get_simpset sg = #simpset (get_sel_upd sg);
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   437
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   438
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   439
fun put_sel_upd names simps thy =
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   440
  let
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   441
    val sels = map (rpair ()) names;
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   442
    val upds = map (suffix updateN) names ~~ names;
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   443
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   444
    val {records, sel_upd = {selectors, updates, simpset}, field_splits} = RecordsData.get thy;
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   445
    val data = make_record_data records
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   446
      {selectors = Symtab.extend (selectors, sels),
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   447
        updates = Symtab.extend (updates, upds),
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   448
        simpset = Simplifier.addsimps (simpset, simps)}
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   449
      field_splits;
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   450
  in RecordsData.put data thy end;
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   451
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   452
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   453
(* access 'field_splits' *)
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   454
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   455
fun add_record_splits splits thy =
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   456
  let
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   457
    val {records, sel_upd, field_splits = {fields, simpset}} = RecordsData.get thy;
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   458
    val flds = map (rpair () o fst) splits;
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   459
    val simps = map snd splits;
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   460
    val data = make_record_data records sel_upd
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   461
      {fields = Symtab.extend (fields, flds), simpset = Simplifier.addsimps (simpset, simps)};
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   462
  in RecordsData.put data thy end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   463
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   464
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   465
(* parent records *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   466
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   467
fun inst_record thy (types, name) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   468
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   469
    val sign = Theory.sign_of thy;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   470
    fun err msg = error (msg ^ " parent record " ^ quote name);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   471
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   472
    val {args, parent, fields, simps, induct, cases} =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   473
      (case get_record thy name of Some info => info | None => err "Unknown");
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   474
    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
   475
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   476
    fun bad_inst ((x, S), T) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   477
      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
   478
    val bads = mapfilter bad_inst (args ~~ types);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   479
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   480
    val inst = map fst args ~~ types;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   481
    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
   482
  in
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   483
    if not (null bads) then
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   484
      err ("Ill-sorted instantiation of " ^ commas bads ^ " in")
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   485
    else (apsome (apfst (map subst)) parent, map (apsnd subst) fields, simps, induct, cases)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   486
  end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   487
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   488
fun add_parents thy (None, parents) = parents
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   489
  | add_parents thy (Some (types, name), parents) =
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   490
      let val (parent, fields, simps, induct, cases) = inst_record thy (types, name)
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   491
      in add_parents thy (parent, make_parent_info name fields simps induct cases :: parents) end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   492
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   493
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   494
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   495
(** record simproc **)
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   496
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   497
local
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   498
8100
6186ee807f2e replaced HOLogic.termTVar by HOLogic.termT;
wenzelm
parents: 7178
diff changeset
   499
val sel_upd_pat = [Thm.read_cterm (Theory.sign_of HOL.thy) ("s (u k r)", HOLogic.termT)];
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   500
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   501
fun proc sg _ t =
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   502
  (case t of (sel as Const (s, _)) $ ((upd as Const (u, _)) $ k $ r) =>
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   503
    (case get_selectors sg s of Some () =>
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   504
      (case get_updates sg u of Some u_name =>
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   505
        let
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   506
          fun atomize x t = Free (x, fastype_of t);
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   507
          val k' = atomize "k" k;
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   508
          val r' = atomize "r" r;
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   509
          val t' = sel $ (upd $ k' $ r');
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   510
          val prove = mk_meta_eq o prove_simp sg (get_simpset sg) [] [];
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   511
        in
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   512
          if u_name = s then Some (prove (t' === k'))
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   513
          else Some (prove (t' === sel $ r'))
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   514
        end
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   515
      | None => None)
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   516
    | None => None)
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   517
  | _ => None);
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   518
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   519
in
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   520
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   521
val record_simproc = Simplifier.mk_simproc "record_simp" sel_upd_pat proc;
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   522
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   523
end;
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   524
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   525
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   526
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   527
(** record field splitting **)
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   528
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   529
(* tactic *)
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   530
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   531
fun record_split_tac i st =
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   532
  let
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   533
    val {field_splits = {fields, simpset}, ...} = RecordsData.get_sg (Thm.sign_of_thm st);
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   534
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   535
    fun is_fieldT (_, Type (a, [_, _])) = is_some (Symtab.lookup (fields, a))
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   536
      | is_fieldT _ = false;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   537
    val params = Logic.strip_params (Library.nth_elem (i - 1, Thm.prems_of st));
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   538
  in
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   539
    if exists is_fieldT params then Simplifier.full_simp_tac simpset i st
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   540
    else Seq.empty
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   541
  end handle Library.LIST _ => Seq.empty;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   542
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   543
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   544
(* wrapper *)
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   545
5707
b0e631634b5a field_injects [iffs];
wenzelm
parents: 5698
diff changeset
   546
val record_split_name = "record_split_tac";
b0e631634b5a field_injects [iffs];
wenzelm
parents: 5698
diff changeset
   547
val record_split_wrapper = (record_split_name, fn tac => record_split_tac ORELSE' tac);
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   548
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   549
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   550
(* method *)
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   551
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   552
val record_split_method =
9705
8eecca293907 Method.SIMPLE_METHOD';
wenzelm
parents: 9626
diff changeset
   553
  ("record_split", Method.no_args (Method.SIMPLE_METHOD' HEADGOAL record_split_tac),
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   554
    "split record fields");
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   555
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
   556
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   557
(** internal theory extenders **)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   558
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   559
(* field_typedefs *)
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   560
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   561
fun field_typedefs zeta moreT names theory =
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   562
  let
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   563
    val alpha = "'a";
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   564
    val aT = TFree (alpha, HOLogic.termS);
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   565
    val UNIV = HOLogic.mk_UNIV (HOLogic.mk_prodT (aT, moreT));
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   566
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   567
    fun type_def (thy, name) =
11940
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   568
      let val (thy', {type_definition, set_def = Some def, ...}) =
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   569
        thy |> setmp TypedefPackage.quiet_mode true
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   570
          (TypedefPackage.add_typedef_i true None
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   571
            (suffix field_typeN (Sign.base_name name), [alpha, zeta], Syntax.NoSyn) UNIV None
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   572
          (Tactic.rtac UNIV_witness 1))
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   573
      in (thy', Tactic.rewrite_rule [def] type_definition) end
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   574
  in foldl_map type_def (theory, names) end;
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   575
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   576
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   577
(* field_definitions *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   578
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   579
fun field_definitions fields names xs alphas zeta moreT more vars named_vars thy =
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   580
  let
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   581
    val sign = Theory.sign_of thy;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   582
    val base = Sign.base_name;
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   583
    val full_path = Sign.full_name_path sign;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   584
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   585
    val xT = TFree (variant alphas "'x", HOLogic.termS);
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   586
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   587
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   588
    (* prepare declarations and definitions *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   589
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   590
    (*field constructors*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   591
    val field_decls = map (mk_fieldC moreT) fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   592
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   593
    fun mk_field_spec ((c, T), v) =
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   594
      Term.head_of (mk_field ((c, v), more)) :==
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   595
        lambda v (lambda more (mk_Abs moreT (c, T) $ (HOLogic.mk_prod (v, more))));
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   596
    val field_specs = map mk_field_spec (fields ~~ vars);
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   597
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   598
    (*field destructors*)
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   599
    val dest_decls = map (mk_fstC moreT) fields @ map (mk_sndC moreT) fields;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   600
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   601
    fun mk_dest_spec dest sel (c, T) =
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   602
      let val p = Free ("p", mk_fieldT ((c, T), moreT));
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   603
      in Term.head_of (dest p) :== lambda p (sel (mk_Rep moreT (c, T) $ p)) end;
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   604
    val dest_specs1 = map (mk_dest_spec mk_fst HOLogic.mk_fst) fields;
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   605
    val dest_specs2 = map (mk_dest_spec mk_snd HOLogic.mk_snd) fields;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   606
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   607
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   608
    (* 1st stage: defs_thy *)
5713
27d4fcf5fe5f fixed field_injects;
wenzelm
parents: 5707
diff changeset
   609
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   610
    val (defs_thy, (((typedefs, field_defs), dest_defs1), dest_defs2)) =
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   611
      thy
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   612
      |> field_typedefs zeta moreT names
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   613
      |>> (Theory.add_consts_i o map (Syntax.no_syn o apfst base)) (field_decls @ dest_decls)
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   614
      |>>> (PureThy.add_defs_i false o map Thm.no_attributes) field_specs
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   615
      |>>> (PureThy.add_defs_i false o map Thm.no_attributes) dest_specs1
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   616
      |>>> (PureThy.add_defs_i false o map Thm.no_attributes) dest_specs2;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   617
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   618
    val prod_types = map (fn (((a, b), c), d) => product_typeI OF [a, b, c, d])
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   619
      (typedefs ~~ field_defs ~~ dest_defs1 ~~ dest_defs2);
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   620
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   621
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   622
    (* 2nd stage: thms_thy *)
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   623
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   624
    fun make th = map (fn prod_type => Drule.standard (th OF [prod_type])) prod_types;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   625
11940
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   626
    val dest_convs = make product_type_conv1 @ make product_type_conv2;
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   627
    val field_injects = make product_type_inject;
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   628
    val field_inducts = make product_type_induct;
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   629
    val field_cases = make product_type_cases;
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   630
    val field_splits = make product_type_split_paired_all;
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   631
11940
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   632
    val (thms_thy, [field_defs', dest_defs', dest_convs', field_injects',
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   633
        field_splits', field_inducts', field_cases']) = defs_thy
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   634
      |> (PureThy.add_thmss o map Thm.no_attributes)
11940
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   635
       [("field_defs", field_defs),
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   636
        ("dest_defs", dest_defs1 @ dest_defs2),
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   637
        ("dest_convs", dest_convs),
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   638
        ("field_injects", field_injects),
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   639
        ("field_splits", field_splits),
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   640
        ("field_inducts", field_inducts),
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   641
        ("field_cases", field_cases)];
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   642
11940
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   643
  in (thms_thy, dest_convs', field_injects', field_splits', field_inducts', field_cases') end;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   644
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   645
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   646
(* record_definition *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   647
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   648
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
   649
  let
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   650
    val sign = Theory.sign_of thy;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   651
    val full = Sign.full_name_path sign bname;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   652
    val base = Sign.base_name;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   653
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   654
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   655
    (* basic components *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   656
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   657
    val alphas = map fst args;
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   658
    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
   659
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   660
    val previous = if null parents then None else Some (last_elem parents);
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   661
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   662
    val parent_fields = flat (map #fields parents);
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   663
    val parent_names = map fst parent_fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   664
    val parent_types = map snd parent_fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   665
    val parent_len = length parent_fields;
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   666
    val parent_xs = variantlist (map (base o fst) parent_fields, [moreN, rN]);
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   667
    val parent_vars = ListPair.map Free (parent_xs, parent_types);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   668
    val parent_named_vars = parent_names ~~ parent_vars;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   669
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   670
    val fields = map (apfst full) bfields;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   671
    val names = map fst fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   672
    val types = map snd fields;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   673
    val len = length fields;
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   674
    val xs = variantlist (map fst bfields, moreN :: rN :: parent_xs);
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   675
    val vars = ListPair.map Free (xs, types);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   676
    val named_vars = names ~~ vars;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   677
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   678
    val all_fields = parent_fields @ fields;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   679
    val all_names = parent_names @ names;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   680
    val all_types = parent_types @ types;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   681
    val all_len = parent_len + len;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   682
    val all_xs = parent_xs @ xs;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   683
    val all_vars = parent_vars @ vars;
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   684
    val all_named_vars = parent_named_vars @ named_vars;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   685
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   686
    val zeta = variant alphas "'z";
11940
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   687
    val moreT = TFree (zeta, HOLogic.termS);
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   688
    val more = Free (moreN, moreT);
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   689
    val full_moreN = full moreN;
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   690
    fun more_part t = mk_more t full_moreN;
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   691
    fun more_part_update t x = mk_more_update t (full_moreN, x);
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   692
    val all_types_more = all_types @ [moreT];
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   693
    val all_xs_more = all_xs @ [moreN];
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   694
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   695
    val parent_more = funpow parent_len mk_snd;
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   696
    val idxs = 0 upto (len - 1);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   697
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   698
    val parentT = if null parent_fields then [] else [mk_recordT (parent_fields, HOLogic.unitT)];
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   699
    val r_parent = if null parent_fields then [] else [Free (rN, hd parentT)];
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   700
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   701
    val rec_schemeT = mk_recordT (all_fields, moreT);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   702
    val rec_scheme = mk_record (all_named_vars, more);
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   703
    val recT = mk_recordT (all_fields, HOLogic.unitT);
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   704
    val rec_ = mk_record (all_named_vars, HOLogic.unit);
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   705
    val r_scheme = Free (rN, rec_schemeT);
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   706
    val r = Free (rN, recT);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   707
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   708
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   709
    (* prepare print translation functions *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   710
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   711
    val field_tr's =
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   712
      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
   713
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   714
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   715
    (* prepare declarations *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   716
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   717
    val sel_decls = map (mk_selC rec_schemeT) bfields @
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   718
      [mk_moreC rec_schemeT (moreN, moreT)];
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   719
    val update_decls = map (mk_updateC rec_schemeT) bfields @
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   720
      [mk_more_updateC rec_schemeT (moreN, moreT)];
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   721
    val make_decl = (makeN, parentT ---> types ---> recT);
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   722
    val extend_decl = (extendN, recT --> moreT --> rec_schemeT);
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   723
    val truncate_decl = (truncateN, rec_schemeT --> recT);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   724
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   725
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   726
    (* prepare definitions *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   727
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   728
    (*record (scheme) type abbreviation*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   729
    val recordT_specs =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   730
      [(suffix schemeN bname, alphas @ [zeta], rec_schemeT, Syntax.NoSyn),
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   731
        (bname, alphas, recT, Syntax.NoSyn)];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   732
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   733
    (*selectors*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   734
    fun mk_sel_spec (i, c) =
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   735
      mk_sel r_scheme c :== mk_fst (funpow i mk_snd (parent_more r_scheme));
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   736
    val sel_specs =
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   737
      ListPair.map mk_sel_spec (idxs, names) @
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   738
        [more_part r_scheme :== funpow len mk_snd (parent_more r_scheme)];
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   739
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   740
    (*updates*)
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   741
    val all_sels = mk_named_sels all_names r_scheme;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   742
    fun mk_upd_spec (i, (c, x)) =
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   743
      mk_update r_scheme (c, x) :==
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   744
        mk_record (nth_update (c, x) (parent_len + i, all_sels), more_part r_scheme)
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   745
    val update_specs =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   746
      ListPair.map mk_upd_spec (idxs, named_vars) @
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   747
        [more_part_update r_scheme more :== mk_record (all_sels, more)];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   748
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   749
    (*derived operations*)
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   750
    val make_spec = Const (full makeN, parentT ---> types ---> recT) $$ r_parent $$ vars :==
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   751
      mk_record (flat (map (mk_named_sels parent_names) r_parent) @ named_vars, HOLogic.unit);
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   752
    val extend_spec = Const (full extendN, recT --> moreT --> rec_schemeT) $ r $ more :==
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   753
      mk_record (mk_named_sels all_names r, more);
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   754
    val truncate_spec = Const (full truncateN, rec_schemeT --> recT) $ r_scheme :==
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   755
      mk_record (all_sels, HOLogic.unit);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   756
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   757
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   758
    (* prepare propositions *)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   759
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   760
    (*selectors*)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   761
    val sel_props =
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   762
      map (fn (c, x) => mk_sel rec_scheme c === x) named_vars @
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   763
        [more_part rec_scheme === more];
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   764
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   765
    (*updates*)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   766
    fun mk_upd_prop (i, (c, T)) =
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   767
      let val x' = Free (variant all_xs (base c ^ "'"), T) in
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   768
        mk_update rec_scheme (c, x') ===
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   769
          mk_record (nth_update (c, x') (parent_len + i, all_named_vars), more)
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   770
      end;
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   771
    val update_props =
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   772
      ListPair.map mk_upd_prop (idxs, fields) @
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   773
        let val more' = Free (variant all_xs (moreN ^ "'"), moreT)
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   774
        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
   775
9626
c4a45149cc46 sel_upd proc: include 'more' pseudo-field;
wenzelm
parents: 9315
diff changeset
   776
    (*equality*)
c4a45149cc46 sel_upd proc: include 'more' pseudo-field;
wenzelm
parents: 9315
diff changeset
   777
    fun mk_sel_eq (t, T) =
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   778
      let val t' = Term.abstract_over (r_scheme, t)
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   779
      in Trueprop (HOLogic.eq_const T $ Term.incr_boundvars 1 t' $ t') end;
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   780
    val sel_eqs =
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   781
      map2 mk_sel_eq (map (mk_sel r_scheme) all_names @ [more_part r_scheme], all_types @ [moreT]);
9626
c4a45149cc46 sel_upd proc: include 'more' pseudo-field;
wenzelm
parents: 9315
diff changeset
   782
    val equality_prop =
c4a45149cc46 sel_upd proc: include 'more' pseudo-field;
wenzelm
parents: 9315
diff changeset
   783
      Term.all rec_schemeT $ (Abs ("r", rec_schemeT,
c4a45149cc46 sel_upd proc: include 'more' pseudo-field;
wenzelm
parents: 9315
diff changeset
   784
        Term.all rec_schemeT $ (Abs ("r'", rec_schemeT,
c4a45149cc46 sel_upd proc: include 'more' pseudo-field;
wenzelm
parents: 9315
diff changeset
   785
          Logic.list_implies (sel_eqs,
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   786
            Trueprop (HOLogic.eq_const rec_schemeT $ Bound 1 $ Bound 0))))));
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   787
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   788
    (*induct*)
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   789
    val P = Free ("P", rec_schemeT --> HOLogic.boolT);
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   790
    val P' = Free ("P", recT --> HOLogic.boolT);
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   791
    val induct_scheme_prop =
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   792
      All (all_xs_more ~~ all_types_more) (Trueprop (P $ rec_scheme)) ==> Trueprop (P $ r_scheme);
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   793
    val induct_prop = All (all_xs ~~ all_types) (Trueprop (P' $ rec_)) ==> Trueprop (P' $ r);
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   794
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   795
    (*cases*)
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   796
    val C = Trueprop (Free (variant all_xs_more "C", HOLogic.boolT));
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   797
    val cases_scheme_prop =
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   798
      All (all_xs_more ~~ all_types_more) ((r_scheme === rec_scheme) ==> C) ==> C;
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   799
    val cases_prop = All (all_xs ~~ all_types) ((r === rec_) ==> C) ==> C;
9626
c4a45149cc46 sel_upd proc: include 'more' pseudo-field;
wenzelm
parents: 9315
diff changeset
   800
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   801
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   802
    (* 1st stage: fields_thy *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   803
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   804
    val (fields_thy, field_simps, field_injects, field_splits, field_inducts, field_cases) =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   805
      thy
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   806
      |> Theory.add_path bname
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   807
      |> field_definitions fields names xs alphas zeta moreT more vars named_vars;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   808
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5890
diff changeset
   809
    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
   810
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   811
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   812
    (* 2nd stage: defs_thy *)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   813
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   814
    val (defs_thy, (((sel_defs, update_defs), derived_defs))) =
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   815
      fields_thy
9626
c4a45149cc46 sel_upd proc: include 'more' pseudo-field;
wenzelm
parents: 9315
diff changeset
   816
      |> add_record_splits named_splits
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   817
      |> Theory.parent_path
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   818
      |> 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
   819
      |> Theory.add_path bname
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   820
      |> Theory.add_trfuns ([], [], field_tr's, [])
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   821
      |> (Theory.add_consts_i o map Syntax.no_syn)
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   822
        (sel_decls @ update_decls @ [make_decl, extend_decl, truncate_decl])
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   823
      |> (PureThy.add_defs_i false o map Thm.no_attributes) sel_specs
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   824
      |>>> (PureThy.add_defs_i false o map Thm.no_attributes) update_specs
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   825
      |>>> (PureThy.add_defs_i false o map Thm.no_attributes)
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   826
        [make_spec, extend_spec, truncate_spec];
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   827
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   828
    val defs_sg = Theory.sign_of defs_thy;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   829
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   830
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   831
    (* 3rd stage: thms_thy *)
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   832
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   833
    val parent_simps = flat (map #simps parents);
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   834
    val prove = prove_simp defs_sg HOL_basic_ss [];
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   835
    val prove' = prove_simp defs_sg HOL_ss;
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   836
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   837
    val sel_convs = map (prove (parent_simps @ sel_defs @ field_simps)) sel_props;
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   838
    val update_convs = map (prove (parent_simps @ update_defs @ sel_convs)) update_props;
9626
c4a45149cc46 sel_upd proc: include 'more' pseudo-field;
wenzelm
parents: 9315
diff changeset
   839
    val equality =
c4a45149cc46 sel_upd proc: include 'more' pseudo-field;
wenzelm
parents: 9315
diff changeset
   840
      prove' [ALLGOALS record_split_tac] (parent_simps @ sel_convs @ field_injects) equality_prop;
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   841
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   842
    val induct_scheme = prove_goal defs_sg induct_scheme_prop (fn prems =>
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   843
        (case previous of Some {induct, ...} => res_inst_tac [(rN, rN)] induct 1
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   844
        | None => all_tac) :: map (fn rule => try_param_tac "p" rN rule 1) field_inducts @
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   845
        [resolve_tac prems 1]);
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   846
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   847
    val induct = prove_goal defs_sg induct_prop (fn prems =>
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   848
        [res_inst_tac [(rN, rN)] induct_scheme 1,
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   849
         try_param_tac "x" "more" unit_induct 1, resolve_tac prems 1]);
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   850
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   851
    val cases_scheme = prove_goal defs_sg cases_scheme_prop (fn prems =>
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   852
        Method.insert_tac prems 1 ::
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   853
        (case previous of Some {cases, ...} => try_param_tac rN rN cases 1
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   854
        | None => all_tac) :: map (fn rule => try_param_tac "p" rN rule 1) field_cases @
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   855
        [Simplifier.asm_full_simp_tac HOL_basic_ss 1]);
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   856
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   857
    val cases = prove_goal defs_sg cases_prop (fn prems =>
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   858
        [Method.insert_tac prems 1, res_inst_tac [(rN, rN)] cases_scheme 1,
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   859
         Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps [unit_all_eq1]) 1]);
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   860
11940
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   861
    val (thms_thy, ([sel_convs', update_convs', sel_defs', update_defs', _],
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   862
        [equality', induct_scheme', induct', cases_scheme', cases'])) =
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   863
      defs_thy
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   864
      |> (PureThy.add_thmss o map Thm.no_attributes)
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   865
       [("select_convs", sel_convs),
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   866
        ("update_convs", update_convs),
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   867
        ("select_defs", sel_defs),
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   868
        ("update_defs", update_defs),
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   869
        ("derived_defs", derived_defs)]
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   870
      |>>> PureThy.add_thms
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   871
       [(("equality", equality), [Classical.xtra_intro_global]),
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   872
        (("induct_scheme", induct_scheme), [RuleCases.case_names [fieldsN],
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   873
          InductAttrib.induct_type_global (suffix schemeN name)]),
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   874
        (("induct", induct), [RuleCases.case_names [fieldsN],
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   875
          InductAttrib.induct_type_global name]),
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   876
        (("cases_scheme", cases_scheme), [RuleCases.case_names [fieldsN],
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   877
          InductAttrib.cases_type_global (suffix schemeN name)]),
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   878
        (("cases", cases), [RuleCases.case_names [fieldsN],
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   879
          InductAttrib.cases_type_global name])];
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   880
11959
ed914e8a0fd1 exclude field_simps from user-level "simps";
wenzelm
parents: 11940
diff changeset
   881
    val simps = sel_convs' @ update_convs' @ [equality'];
6519
5bd1c469e742 iff_add_global (from simpdata.ML);
wenzelm
parents: 6394
diff changeset
   882
    val iffs = field_injects;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   883
11940
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   884
    val thms_thy' =
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   885
      thms_thy |> (#1 oo PureThy.add_thmss)
5707
b0e631634b5a field_injects [iffs];
wenzelm
parents: 5698
diff changeset
   886
        [(("simps", simps), [Simplifier.simp_add_global]),
6519
5bd1c469e742 iff_add_global (from simpdata.ML);
wenzelm
parents: 6394
diff changeset
   887
         (("iffs", iffs), [iff_add_global])];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   888
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   889
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   890
    (* 4th stage: final_thy *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   891
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   892
    val final_thy =
11940
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   893
      thms_thy'
11959
ed914e8a0fd1 exclude field_simps from user-level "simps";
wenzelm
parents: 11940
diff changeset
   894
      |> put_record name (make_record_info args parent fields (field_simps @ simps)
ed914e8a0fd1 exclude field_simps from user-level "simps";
wenzelm
parents: 11940
diff changeset
   895
          induct_scheme' cases_scheme')
11940
80365073b8b3 removed "more" sort;
wenzelm
parents: 11934
diff changeset
   896
      |> put_sel_upd (names @ [full_moreN]) (field_simps @ sel_defs' @ update_defs')
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   897
      |> Theory.parent_path;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   898
6519
5bd1c469e742 iff_add_global (from simpdata.ML);
wenzelm
parents: 6394
diff changeset
   899
  in (final_thy, {simps = simps, iffs = iffs}) end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   900
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   901
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   902
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   903
(** theory extender interface **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   904
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   905
(* prepare arguments *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   906
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   907
(*note: read_raw_typ avoids expanding type abbreviations*)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   908
fun read_raw_parent sign s =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   909
  (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
   910
    Type (name, Ts) => (Ts, name)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   911
  | _ => error ("Bad parent record specification: " ^ quote s));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   912
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   913
fun read_typ sign (env, s) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   914
  let
5060
7b86df67cc1a def_sort;
wenzelm
parents: 5052
diff changeset
   915
    fun def_sort (x, ~1) = assoc (env, x)
7b86df67cc1a def_sort;
wenzelm
parents: 5052
diff changeset
   916
      | def_sort _ = None;
7b86df67cc1a def_sort;
wenzelm
parents: 5052
diff changeset
   917
    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
   918
  in (Term.add_typ_tfrees (T, env), T) end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   919
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   920
fun cert_typ sign (env, raw_T) =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   921
  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
   922
  in (Term.add_typ_tfrees (T, env), T) end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   923
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   924
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   925
(* add_record *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   926
4895
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   927
(*we do all preparations and error checks here, deferring the real
0fad2acb45fb misc tuning;
wenzelm
parents: 4894
diff changeset
   928
  work to record_definition*)
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   929
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   930
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
   931
  let
4970
8b65444edbb0 Changed require to requires for MLWorks
paulson
parents: 4967
diff changeset
   932
    val _ = Theory.requires thy "Record" "record definitions";
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   933
    val sign = Theory.sign_of thy;
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
   934
    val _ = message ("Defining record " ^ quote bname ^ " ...");
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   935
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   936
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   937
    (* parents *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   938
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   939
    fun prep_inst T = snd (cert_typ sign ([], T));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   940
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   941
    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
   942
      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
   943
    val parents = add_parents thy (parent, []);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   944
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   945
    val init_env =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   946
      (case parent of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   947
        None => []
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   948
      | Some (types, _) => foldr Term.add_typ_tfrees (types, []));
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   949
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   950
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   951
    (* fields *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   952
4967
c9c481bc0216 foldl_map prep_field;
wenzelm
parents: 4895
diff changeset
   953
    fun prep_field (env, (c, raw_T)) =
c9c481bc0216 foldl_map prep_field;
wenzelm
parents: 4895
diff changeset
   954
      let val (env', T) = prep_typ sign (env, raw_T) handle ERROR =>
c9c481bc0216 foldl_map prep_field;
wenzelm
parents: 4895
diff changeset
   955
        error ("The error(s) above occured in field " ^ quote c)
c9c481bc0216 foldl_map prep_field;
wenzelm
parents: 4895
diff changeset
   956
      in (env', (c, T)) end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   957
4967
c9c481bc0216 foldl_map prep_field;
wenzelm
parents: 4895
diff changeset
   958
    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
   959
    val envir_names = map fst envir;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   960
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   961
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   962
    (* args *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   963
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   964
    val defaultS = Sign.defaultS sign;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   965
    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
   966
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   967
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   968
    (* errors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   969
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   970
    val name = Sign.full_name sign bname;
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   971
    val err_dup_record =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   972
      if is_none (get_record thy name) then []
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   973
      else ["Duplicate definition of record " ^ quote name];
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   974
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   975
    val err_dup_parms =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   976
      (case duplicates params of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   977
        [] => []
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   978
      | dups => ["Duplicate parameter(s) " ^ commas dups]);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   979
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   980
    val err_extra_frees =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   981
      (case gen_rems (op =) (envir_names, params) of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   982
        [] => []
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   983
      | extras => ["Extra free type variable(s) " ^ commas extras]);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   984
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   985
    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
   986
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   987
    val err_dup_fields =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   988
      (case duplicates (map fst bfields) of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   989
        [] => []
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   990
      | dups => ["Duplicate field(s) " ^ commas_quote dups]);
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   991
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   992
    val err_bad_fields =
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   993
      if forall (not_equal moreN o fst) bfields then []
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
   994
      else ["Illegal field name " ^ quote moreN];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   995
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   996
    val err_dup_sorts =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   997
      (case duplicates envir_names of
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   998
        [] => []
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   999
      | dups => ["Inconsistent sort constraints for " ^ commas dups]);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1000
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1001
    val errs =
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
  1002
      err_dup_record @ err_dup_parms @ err_extra_frees @ err_no_fields @
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
  1003
      err_dup_fields @ err_bad_fields @ err_dup_sorts;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1004
  in
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
  1005
    if null errs then () else error (cat_lines errs);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1006
    thy |> record_definition (args, bname) parent parents bfields
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1007
  end
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1008
  handle ERROR => error ("Failed to define record " ^ quote bname);
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1009
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1010
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
  1011
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
  1012
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1013
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1014
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  1015
(** package setup **)
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  1016
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  1017
(* setup theory *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1018
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1019
val setup =
5006
cdc86a914e63 adapted to new theory data interface;
wenzelm
parents: 5001
diff changeset
  1020
 [RecordsData.init,
11490
f9ae28f55178 field_name_ast_tr superceded by constify_ast_tr in Pure;
wenzelm
parents: 11473
diff changeset
  1021
  Theory.add_trfuns ([], parse_translation, [], []),
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  1022
  Method.add_methods [record_split_method],
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
  1023
  Simplifier.change_simpset_of Simplifier.addsimprocs [record_simproc]];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1024
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1025
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  1026
(* outer syntax *)
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  1027
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
  1028
local structure P = OuterParse and K = OuterSyntax.Keyword in
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  1029
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  1030
val record_decl =
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
  1031
  P.type_args -- P.name -- (P.$$$ "=" |-- Scan.option (P.typ --| P.$$$ "+")
6729
b6e167580a32 formal comments (still dummy);
wenzelm
parents: 6723
diff changeset
  1032
    -- Scan.repeat1 (P.name -- (P.$$$ "::" |-- P.typ) --| P.marg_comment));
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  1033
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  1034
val recordP =
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
  1035
  OuterSyntax.command "record" "define extensible record" K.thy_decl
6519
5bd1c469e742 iff_add_global (from simpdata.ML);
wenzelm
parents: 6394
diff changeset
  1036
    (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
  1037
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  1038
val _ = OuterSyntax.add_parsers [recordP];
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  1039
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1040
end;
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
  1041
6384
eed1273c9146 local open OuterParse;
wenzelm
parents: 6358
diff changeset
  1042
end;
eed1273c9146 local open OuterParse;
wenzelm
parents: 6358
diff changeset
  1043
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
  1044
structure BasicRecordPackage: BASIC_RECORD_PACKAGE = RecordPackage;
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
  1045
open BasicRecordPackage;