src/HOL/Tools/record.ML
author boehmes
Mon, 08 Nov 2010 12:13:44 +0100
changeset 40424 7550b2cba1cb
parent 40315 11846d9800de
child 40722 441260986b63
permissions -rw-r--r--
better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31136
diff changeset
     1
(*  Title:      HOL/Tools/record.ML
32763
ebfaf9e3c03a tuned header;
wenzelm
parents: 32761
diff changeset
     2
    Author:     Wolfgang Naraschewski, TU Muenchen
ebfaf9e3c03a tuned header;
wenzelm
parents: 32761
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
ebfaf9e3c03a tuned header;
wenzelm
parents: 32761
diff changeset
     4
    Author:     Norbert Schirmer, TU Muenchen
ebfaf9e3c03a tuned header;
wenzelm
parents: 32761
diff changeset
     5
    Author:     Thomas Sewell, NICTA
ebfaf9e3c03a tuned header;
wenzelm
parents: 32761
diff changeset
     6
ebfaf9e3c03a tuned header;
wenzelm
parents: 32761
diff changeset
     7
Extensible records with structural subtyping.
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
     8
*)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
     9
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    10
signature RECORD =
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
    11
sig
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    12
  val print_type_abbr: bool Unsynchronized.ref
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    13
  val print_type_as_fields: bool Unsynchronized.ref
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    14
  val timing: bool Unsynchronized.ref
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    15
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    16
  type info =
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
    17
   {args: (string * sort) list,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
    18
    parent: (typ list * string) option,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
    19
    fields: (string * typ) list,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
    20
    extension: (string * typ list),
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
    21
    ext_induct: thm, ext_inject: thm, ext_surjective: thm, ext_split: thm, ext_def: thm,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
    22
    select_convs: thm list, update_convs: thm list, select_defs: thm list, update_defs: thm list,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
    23
    fold_congs: thm list, unfold_congs: thm list, splits: thm list, defs: thm list,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
    24
    surjective: thm, equality: thm, induct_scheme: thm, induct: thm, cases_scheme: thm,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
    25
    cases: thm, simps: thm list, iffs: thm list}
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    26
  val get_info: theory -> string -> info option
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    27
  val the_info: theory -> string -> info
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    28
  val add_record: bool -> (string * sort) list * binding -> (typ list * string) option ->
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    29
    (binding * typ * mixfix) list -> theory -> theory
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    30
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
    31
  val last_extT: typ -> (string * typ list) option
32972
45ba8b02e1e4 misc tuning and simplification;
wenzelm
parents: 32970
diff changeset
    32
  val dest_recTs: typ -> (string * typ list) list
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
    33
  val get_extT_fields: theory -> typ -> (string * typ) list * (string * typ)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
    34
  val get_recT_fields: theory -> typ -> (string * typ) list * (string * typ)
26088
9b48d0264ffd added get_parent (for AWE);
wenzelm
parents: 26065
diff changeset
    35
  val get_parent: theory -> string -> (typ list * string) option
9b48d0264ffd added get_parent (for AWE);
wenzelm
parents: 26065
diff changeset
    36
  val get_extension: theory -> string -> (string * typ list) option
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16379
diff changeset
    37
  val get_extinjects: theory -> thm list
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16379
diff changeset
    38
  val get_simpset: theory -> simpset
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    39
  val simproc: simproc
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    40
  val eq_simproc: simproc
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    41
  val upd_simproc: simproc
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    42
  val split_simproc: (term -> int) -> simproc
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    43
  val ex_sel_eq_simproc: simproc
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    44
  val split_tac: int -> tactic
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    45
  val split_simp_tac: thm list -> (term -> int) -> int -> tactic
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    46
  val split_wrapper: string * wrapper
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    47
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    48
  val updateN: string
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    49
  val ext_typeN: string
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
    50
  val extN: string
27278
129574589734 export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27239
diff changeset
    51
  val read_typ: Proof.context -> string -> (string * sort) list -> typ * (string * sort) list
129574589734 export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27239
diff changeset
    52
  val cert_typ: Proof.context -> typ -> (string * sort) list -> typ * (string * sort) list
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18688
diff changeset
    53
  val setup: theory -> theory
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    54
end;
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
    55
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
    56
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
    57
signature ISO_TUPLE_SUPPORT =
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
    58
sig
37470
fcc768dc9dd0 more binding; avoid arcane Rep and Abs prefixes
haftmann
parents: 37469
diff changeset
    59
  val add_iso_tuple_type: binding * (string * sort) list ->
36151
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
    60
    typ * typ -> theory -> (term * term) * theory
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
    61
  val mk_cons_tuple: term * term -> term
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
    62
  val dest_cons_tuple: term -> term * term
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
    63
  val iso_tuple_intros_tac: int -> tactic
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
    64
  val named_cterm_instantiate: (string * cterm) list -> thm -> thm
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
    65
end;
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
    66
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
    67
structure Iso_Tuple_Support: ISO_TUPLE_SUPPORT =
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
    68
struct
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
    69
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
    70
val isoN = "_Tuple_Iso";
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
    71
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
    72
val iso_tuple_intro = @{thm isomorphic_tuple_intro};
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
    73
val iso_tuple_intros = Tactic.build_net @{thms isomorphic_tuple.intros};
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
    74
37470
fcc768dc9dd0 more binding; avoid arcane Rep and Abs prefixes
haftmann
parents: 37469
diff changeset
    75
val tuple_iso_tuple = (@{const_name Record.tuple_iso_tuple}, @{thm tuple_iso_tuple});
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
    76
38252
175a5b4b2c94 tuned comments;
wenzelm
parents: 38012
diff changeset
    77
fun named_cterm_instantiate values thm =  (* FIXME eliminate *)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
    78
  let
35135
1667fd3b051d tuned errors;
wenzelm
parents: 35133
diff changeset
    79
    fun match name ((name', _), _) = name = name';
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
    80
    fun getvar name =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
    81
      (case find_first (match name) (Term.add_vars (prop_of thm) []) of
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
    82
        SOME var => cterm_of (theory_of_thm thm) (Var var)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
    83
      | NONE => raise THM ("named_cterm_instantiate: " ^ name, 0, [thm]));
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
    84
  in
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
    85
    cterm_instantiate (map (apfst getvar) values) thm
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
    86
  end;
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
    87
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
    88
structure Iso_Tuple_Thms = Theory_Data
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
    89
(
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
    90
  type T = thm Symtab.table;
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
    91
  val empty = Symtab.make [tuple_iso_tuple];
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
    92
  val extend = I;
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33385
diff changeset
    93
  fun merge data = Symtab.merge Thm.eq_thm_prop data;   (* FIXME handle Symtab.DUP ?? *)
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
    94
);
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
    95
38534
d2fffb763a58 tuned code
haftmann
parents: 38533
diff changeset
    96
fun get_typedef_info tyco vs (({ rep_type, Abs_name, Rep_name, ...},
38529
4cc2ca4d6237 formally integrated typecopy layer into record package
haftmann
parents: 38401
diff changeset
    97
    { Abs_inject, Rep_inject, Abs_inverse, Rep_inverse, ... }) : Typedef.info) thy =
4cc2ca4d6237 formally integrated typecopy layer into record package
haftmann
parents: 38401
diff changeset
    98
  let
4cc2ca4d6237 formally integrated typecopy layer into record package
haftmann
parents: 38401
diff changeset
    99
    val exists_thm =
4cc2ca4d6237 formally integrated typecopy layer into record package
haftmann
parents: 38401
diff changeset
   100
      UNIV_I
4cc2ca4d6237 formally integrated typecopy layer into record package
haftmann
parents: 38401
diff changeset
   101
      |> Drule.instantiate' [SOME (ctyp_of thy (Logic.varifyT_global rep_type))] [];
4cc2ca4d6237 formally integrated typecopy layer into record package
haftmann
parents: 38401
diff changeset
   102
    val proj_constr = Abs_inverse OF [exists_thm];
38534
d2fffb763a58 tuned code
haftmann
parents: 38533
diff changeset
   103
    val absT = Type (tyco, map TFree vs);
38529
4cc2ca4d6237 formally integrated typecopy layer into record package
haftmann
parents: 38401
diff changeset
   104
  in
4cc2ca4d6237 formally integrated typecopy layer into record package
haftmann
parents: 38401
diff changeset
   105
    thy
38534
d2fffb763a58 tuned code
haftmann
parents: 38533
diff changeset
   106
    |> pair (tyco, ((Rep_inject, proj_constr), Const (Abs_name, rep_type --> absT), absT))
38529
4cc2ca4d6237 formally integrated typecopy layer into record package
haftmann
parents: 38401
diff changeset
   107
  end
4cc2ca4d6237 formally integrated typecopy layer into record package
haftmann
parents: 38401
diff changeset
   108
38534
d2fffb763a58 tuned code
haftmann
parents: 38533
diff changeset
   109
fun do_typedef raw_tyco repT raw_vs thy =
38529
4cc2ca4d6237 formally integrated typecopy layer into record package
haftmann
parents: 38401
diff changeset
   110
  let
38534
d2fffb763a58 tuned code
haftmann
parents: 38533
diff changeset
   111
    val ctxt = ProofContext.init_global thy |> Variable.declare_typ repT;
38529
4cc2ca4d6237 formally integrated typecopy layer into record package
haftmann
parents: 38401
diff changeset
   112
    val vs = map (ProofContext.check_tfree ctxt) raw_vs;
4cc2ca4d6237 formally integrated typecopy layer into record package
haftmann
parents: 38401
diff changeset
   113
    val tac = Tactic.rtac UNIV_witness 1;
4cc2ca4d6237 formally integrated typecopy layer into record package
haftmann
parents: 38401
diff changeset
   114
  in
4cc2ca4d6237 formally integrated typecopy layer into record package
haftmann
parents: 38401
diff changeset
   115
    thy
4cc2ca4d6237 formally integrated typecopy layer into record package
haftmann
parents: 38401
diff changeset
   116
    |> Typedef.add_typedef_global false (SOME raw_tyco) (raw_tyco, vs, NoSyn)
38534
d2fffb763a58 tuned code
haftmann
parents: 38533
diff changeset
   117
        (HOLogic.mk_UNIV repT) NONE tac
d2fffb763a58 tuned code
haftmann
parents: 38533
diff changeset
   118
    |-> (fn (tyco, info) => get_typedef_info tyco vs info)
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
   119
  end;
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
   120
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   121
fun mk_cons_tuple (left, right) =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   122
  let
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
   123
    val (leftT, rightT) = (fastype_of left, fastype_of right);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   124
    val prodT = HOLogic.mk_prodT (leftT, rightT);
32972
45ba8b02e1e4 misc tuning and simplification;
wenzelm
parents: 32970
diff changeset
   125
    val isomT = Type (@{type_name tuple_isomorphism}, [prodT, leftT, rightT]);
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
   126
  in
37470
fcc768dc9dd0 more binding; avoid arcane Rep and Abs prefixes
haftmann
parents: 37469
diff changeset
   127
    Const (@{const_name Record.iso_tuple_cons}, isomT --> leftT --> rightT --> prodT) $
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
   128
      Const (fst tuple_iso_tuple, isomT) $ left $ right
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
   129
  end;
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
   130
37470
fcc768dc9dd0 more binding; avoid arcane Rep and Abs prefixes
haftmann
parents: 37469
diff changeset
   131
fun dest_cons_tuple (Const (@{const_name Record.iso_tuple_cons}, _) $ Const _ $ t $ u) = (t, u)
32972
45ba8b02e1e4 misc tuning and simplification;
wenzelm
parents: 32970
diff changeset
   132
  | dest_cons_tuple t = raise TERM ("dest_cons_tuple", [t]);
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
   133
37470
fcc768dc9dd0 more binding; avoid arcane Rep and Abs prefixes
haftmann
parents: 37469
diff changeset
   134
fun add_iso_tuple_type (b, alphas) (leftT, rightT) thy =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   135
  let
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   136
    val repT = HOLogic.mk_prodT (leftT, rightT);
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   137
38534
d2fffb763a58 tuned code
haftmann
parents: 38533
diff changeset
   138
    val ((_, ((rep_inject, abs_inverse), absC, absT)), typ_thy) =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   139
      thy
37470
fcc768dc9dd0 more binding; avoid arcane Rep and Abs prefixes
haftmann
parents: 37469
diff changeset
   140
      |> do_typedef b repT alphas
fcc768dc9dd0 more binding; avoid arcane Rep and Abs prefixes
haftmann
parents: 37469
diff changeset
   141
      ||> Sign.add_path (Binding.name_of b); (*FIXME proper prefixing instead*)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   142
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   143
    (*construct a type and body for the isomorphism constant by
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   144
      instantiating the theorem to which the definition will be applied*)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   145
    val intro_inst =
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
   146
      rep_inject RS named_cterm_instantiate [("abst", cterm_of typ_thy absC)] iso_tuple_intro;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   147
    val (_, body) = Logic.dest_equals (List.last (prems_of intro_inst));
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   148
    val isomT = fastype_of body;
37470
fcc768dc9dd0 more binding; avoid arcane Rep and Abs prefixes
haftmann
parents: 37469
diff changeset
   149
    val isom_binding = Binding.suffix_name isoN b;
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
   150
    val isom_name = Sign.full_name typ_thy isom_binding;
32972
45ba8b02e1e4 misc tuning and simplification;
wenzelm
parents: 32970
diff changeset
   151
    val isom = Const (isom_name, isomT);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   152
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   153
    val ([isom_def], cdef_thy) =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   154
      typ_thy
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
   155
      |> Sign.declare_const ((isom_binding, isomT), NoSyn) |> snd
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39134
diff changeset
   156
      |> Global_Theory.add_defs false
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
   157
        [((Binding.conceal (Thm.def_binding isom_binding), Logic.mk_equals (isom, body)), [])];
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   158
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
   159
    val iso_tuple = isom_def RS (abs_inverse RS (rep_inject RS iso_tuple_intro));
37470
fcc768dc9dd0 more binding; avoid arcane Rep and Abs prefixes
haftmann
parents: 37469
diff changeset
   160
    val cons = Const (@{const_name Record.iso_tuple_cons}, isomT --> leftT --> rightT --> absT);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   161
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   162
    val thm_thy =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   163
      cdef_thy
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
   164
      |> Iso_Tuple_Thms.map (Symtab.insert Thm.eq_thm_prop (isom_name, iso_tuple))
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
   165
      |> Sign.restore_naming thy
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   166
  in
32972
45ba8b02e1e4 misc tuning and simplification;
wenzelm
parents: 32970
diff changeset
   167
    ((isom, cons $ isom), thm_thy)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   168
  end;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   169
35133
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   170
val iso_tuple_intros_tac =
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   171
  resolve_from_net_tac iso_tuple_intros THEN'
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   172
  CSUBGOAL (fn (cgoal, i) =>
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   173
    let
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   174
      val thy = Thm.theory_of_cterm cgoal;
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   175
      val goal = Thm.term_of cgoal;
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   176
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
   177
      val isthms = Iso_Tuple_Thms.get thy;
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
   178
      fun err s t = raise TERM ("iso_tuple_intros_tac: " ^ s, [t]);
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   179
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   180
      val goal' = Envir.beta_eta_contract goal;
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   181
      val is =
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   182
        (case goal' of
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   183
          Const (@{const_name Trueprop}, _) $
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   184
            (Const (@{const_name isomorphic_tuple}, _) $ Const is) => is
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   185
        | _ => err "unexpected goal format" goal');
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   186
      val isthm =
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   187
        (case Symtab.lookup isthms (#1 is) of
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   188
          SOME isthm => isthm
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   189
        | NONE => err "no thm found for constant" (Const is));
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   190
    in rtac isthm i end);
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
   191
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
   192
end;
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
   193
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   194
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31136
diff changeset
   195
structure Record: RECORD =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   196
struct
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   197
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   198
val eq_reflection = @{thm eq_reflection};
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   199
val atomize_all = @{thm HOL.atomize_all};
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   200
val atomize_imp = @{thm HOL.atomize_imp};
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   201
val meta_allE = @{thm Pure.meta_allE};
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   202
val prop_subst = @{thm prop_subst};
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   203
val K_record_comp = @{thm K_record_comp};
35136
34206672b168 modernized signature -- proper binding;
wenzelm
parents: 35135
diff changeset
   204
val K_comp_convs = [@{thm o_apply}, K_record_comp];
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   205
val o_assoc = @{thm o_assoc};
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
   206
val id_apply = @{thm id_apply};
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
   207
val id_o_apps = [@{thm id_apply}, @{thm id_o}, @{thm o_id}];
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
   208
val Not_eq_iff = @{thm Not_eq_iff};
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
   209
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   210
val refl_conj_eq = @{thm refl_conj_eq};
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
   211
35133
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   212
val surject_assistI = @{thm iso_tuple_surjective_proof_assistI};
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   213
val surject_assist_idE = @{thm iso_tuple_surjective_proof_assist_idE};
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   214
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   215
val updacc_accessor_eqE = @{thm update_accessor_accessor_eqE};
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   216
val updacc_updator_eqE = @{thm update_accessor_updator_eqE};
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   217
val updacc_eq_idI = @{thm iso_tuple_update_accessor_eq_assist_idI};
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   218
val updacc_eq_triv = @{thm iso_tuple_update_accessor_eq_assist_triv};
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   219
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   220
val updacc_foldE = @{thm update_accessor_congruence_foldE};
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   221
val updacc_unfoldE = @{thm update_accessor_congruence_unfoldE};
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   222
val updacc_noopE = @{thm update_accessor_noopE};
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   223
val updacc_noop_compE = @{thm update_accessor_noop_compE};
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   224
val updacc_cong_idI = @{thm update_accessor_cong_assist_idI};
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   225
val updacc_cong_triv = @{thm update_accessor_cong_assist_triv};
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
   226
val updacc_cong_from_eq = @{thm iso_tuple_update_accessor_cong_from_eq};
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
   227
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   228
val o_eq_dest = @{thm o_eq_dest};
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   229
val o_eq_id_dest = @{thm o_eq_id_dest};
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   230
val o_eq_dest_lhs = @{thm o_eq_dest_lhs};
11832
8fca3665d1ee use abstract product type instead of datatype;
wenzelm
parents: 11739
diff changeset
   231
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   232
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   233
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   234
(** name components **)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   235
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   236
val rN = "r";
15215
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
   237
val wN = "w";
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   238
val moreN = "more";
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   239
val schemeN = "_scheme";
38531
a11a1e4e0403 authentic syntax allows simplification of type names
haftmann
parents: 38530
diff changeset
   240
val ext_typeN = "_ext";
a11a1e4e0403 authentic syntax allows simplification of type names
haftmann
parents: 38530
diff changeset
   241
val inner_typeN = "_inner";
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   242
val extN ="_ext";
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   243
val updateN = "_update";
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   244
val makeN = "make";
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   245
val fields_selN = "fields";
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   246
val extendN = "extend";
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   247
val truncateN = "truncate";
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   248
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
   249
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
   250
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   251
(*** utilities ***)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   252
14709
d01983034ded tuned HOL/record package; enabled record_upd_simproc by default.
schirmer
parents: 14702
diff changeset
   253
fun but_last xs = fst (split_last xs);
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   254
19748
5d05d091eecb fixed bug in type print translations
schirmer
parents: 19343
diff changeset
   255
fun varifyT midx =
5d05d091eecb fixed bug in type print translations
schirmer
parents: 19343
diff changeset
   256
  let fun varify (a, S) = TVar ((a, midx + 1), S);
5d05d091eecb fixed bug in type print translations
schirmer
parents: 19343
diff changeset
   257
  in map_type_tfree varify end;
5d05d091eecb fixed bug in type print translations
schirmer
parents: 19343
diff changeset
   258
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
   259
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
   260
(* timing *)
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
   261
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32335
diff changeset
   262
val timing = Unsynchronized.ref false;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   263
fun timeit_msg s x = if ! timing then (warning s; timeit x) else x ();
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   264
fun timing_msg s = if ! timing then warning s else ();
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
   265
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
   266
12255
wenzelm
parents: 12247
diff changeset
   267
(* syntax *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   268
11927
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   269
val Trueprop = HOLogic.mk_Trueprop;
96f267adc029 provodes induct/cases for use with corresponding Isar methods;
wenzelm
parents: 11923
diff changeset
   270
fun All xs t = Term.list_all_free (xs, t);
4894
32187e0b8b48 'more' selector;
wenzelm
parents: 4890
diff changeset
   271
11934
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   272
infix 0 :== ===;
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   273
infixr 0 ==>;
6c1bf72430b6 derived operations are now: make, extend, truncate (cf. derived_defs);
wenzelm
parents: 11927
diff changeset
   274
37781
2fbbf0a48cef moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents: 37470
diff changeset
   275
val op :== = Misc_Legacy.mk_defpair;
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
   276
val op === = Trueprop o HOLogic.mk_eq;
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
   277
val op ==> = Logic.mk_implies;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   278
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
   279
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   280
(* constructor *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   281
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   282
fun mk_ext (name, T) ts =
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   283
  let val Ts = map fastype_of ts
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
   284
  in list_comb (Const (suffix extN name, Ts ---> T), ts) end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   285
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
   286
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   287
(* selector *)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   288
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   289
fun mk_selC sT (c, T) = (c, sT --> T);
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   290
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   291
fun mk_sel s (c, T) =
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   292
  let val sT = fastype_of s
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   293
  in Const (mk_selC sT (c, T)) $ s end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   294
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
   295
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   296
(* updates *)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   297
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   298
fun mk_updC sfx sT (c, T) = (suffix sfx c, (T --> T) --> sT --> sT);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   299
21226
a607ae87ee81 field-update in records is generalised to take a function on the field
schirmer
parents: 21109
diff changeset
   300
fun mk_upd' sfx c v sT =
a607ae87ee81 field-update in records is generalised to take a function on the field
schirmer
parents: 21109
diff changeset
   301
  let val vT = domain_type (fastype_of v);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   302
  in Const (mk_updC sfx sT (c, vT)) $ v end;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   303
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   304
fun mk_upd sfx c v s = mk_upd' sfx c v (fastype_of s) $ s;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   305
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
   306
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   307
(* types *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   308
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
   309
fun dest_recT (typ as Type (c_ext_type, Ts as (_ :: _))) =
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   310
      (case try (unsuffix ext_typeN) c_ext_type of
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31136
diff changeset
   311
        NONE => raise TYPE ("Record.dest_recT", [typ], [])
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   312
      | SOME c => ((c, Ts), List.last Ts))
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31136
diff changeset
   313
  | dest_recT typ = raise TYPE ("Record.dest_recT", [typ], []);
5197
69c77ed95ba3 added more_update;
wenzelm
parents: 5060
diff changeset
   314
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
   315
val is_recT = can dest_recT;
11833
wenzelm
parents: 11832
diff changeset
   316
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   317
fun dest_recTs T =
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   318
  let val ((c, Ts), U) = dest_recT T
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   319
  in (c, Ts) :: dest_recTs U
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   320
  end handle TYPE _ => [];
14255
e6e3e3f0deed Records:
schirmer
parents: 14079
diff changeset
   321
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   322
fun last_extT T =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   323
  let val ((c, Ts), U) = dest_recT T in
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   324
    (case last_extT U of
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   325
      NONE => SOME (c, Ts)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   326
    | SOME l => SOME l)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   327
  end handle TYPE _ => NONE;
14255
e6e3e3f0deed Records:
schirmer
parents: 14079
diff changeset
   328
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
   329
fun rec_id i T =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   330
  let
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   331
    val rTs = dest_recTs T;
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
   332
    val rTs' = if i < 0 then rTs else take i rTs;
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   333
  in implode (map #1 rTs') end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   334
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
   335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
   336
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   337
(*** extend theory by record definition ***)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   338
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   339
(** record info **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   340
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   341
(* type info and parent_info *)
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   342
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   343
type info =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   344
 {args: (string * sort) list,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   345
  parent: (typ list * string) option,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   346
  fields: (string * typ) list,
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   347
  extension: (string * typ list),
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   348
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   349
  ext_induct: thm,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   350
  ext_inject: thm,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   351
  ext_surjective: thm,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   352
  ext_split: thm,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   353
  ext_def: thm,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   354
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   355
  select_convs: thm list,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   356
  update_convs: thm list,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   357
  select_defs: thm list,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   358
  update_defs: thm list,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   359
  fold_congs: thm list,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   360
  unfold_congs: thm list,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   361
  splits: thm list,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   362
  defs: thm list,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   363
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   364
  surjective: thm,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   365
  equality: thm,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   366
  induct_scheme: thm,
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
   367
  induct: thm,
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   368
  cases_scheme: thm,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   369
  cases: thm,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   370
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   371
  simps: thm list,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   372
  iffs: thm list};
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   373
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   374
fun make_info args parent fields extension
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   375
    ext_induct ext_inject ext_surjective ext_split ext_def
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   376
    select_convs update_convs select_defs update_defs fold_congs unfold_congs splits defs
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   377
    surjective equality induct_scheme induct cases_scheme cases
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   378
    simps iffs : info =
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
   379
 {args = args, parent = parent, fields = fields, extension = extension,
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   380
  ext_induct = ext_induct, ext_inject = ext_inject, ext_surjective = ext_surjective,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   381
  ext_split = ext_split, ext_def = ext_def, select_convs = select_convs,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   382
  update_convs = update_convs, select_defs = select_defs, update_defs = update_defs,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   383
  fold_congs = fold_congs, unfold_congs = unfold_congs, splits = splits, defs = defs,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   384
  surjective = surjective, equality = equality, induct_scheme = induct_scheme,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   385
  induct = induct, cases_scheme = cases_scheme, cases = cases, simps = simps, iffs = iffs};
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   386
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   387
type parent_info =
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   388
 {name: string,
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   389
  fields: (string * typ) list,
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   390
  extension: (string * typ list),
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   391
  induct_scheme: thm,
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   392
  ext_def: thm};
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   393
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   394
fun make_parent_info name fields extension ext_def induct_scheme : parent_info =
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
   395
 {name = name, fields = fields, extension = extension,
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   396
  ext_def = ext_def, induct_scheme = induct_scheme};
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   397
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22747
diff changeset
   398
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22747
diff changeset
   399
(* theory data *)
5001
9de7fda0a6df accomodate tuned version of theory data;
wenzelm
parents: 4970
diff changeset
   400
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   401
type data =
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   402
 {records: info Symtab.table,
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   403
  sel_upd:
32744
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   404
   {selectors: (int * bool) Symtab.table,
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   405
    updates: string Symtab.table,
32744
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   406
    simpset: Simplifier.simpset,
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   407
    defset: Simplifier.simpset,
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   408
    foldcong: Simplifier.simpset,
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   409
    unfoldcong: Simplifier.simpset},
14255
e6e3e3f0deed Records:
schirmer
parents: 14079
diff changeset
   410
  equalities: thm Symtab.table,
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
   411
  extinjects: thm list,
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   412
  extsplit: thm Symtab.table,  (*maps extension name to split rule*)
35135
1667fd3b051d tuned errors;
wenzelm
parents: 35133
diff changeset
   413
  splits: (thm * thm * thm * thm) Symtab.table,  (*!!, ALL, EX - split-equalities, induct rule*)
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   414
  extfields: (string * typ) list Symtab.table,  (*maps extension to its fields*)
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   415
  fieldext: (string * typ list) Symtab.table};  (*maps field to its extension*)
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   416
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   417
fun make_data
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   418
    records sel_upd equalities extinjects extsplit splits extfields fieldext =
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
   419
 {records = records, sel_upd = sel_upd,
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
   420
  equalities = equalities, extinjects=extinjects, extsplit = extsplit, splits = splits,
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   421
  extfields = extfields, fieldext = fieldext }: data;
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   422
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   423
structure Data = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22747
diff changeset
   424
(
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   425
  type T = data;
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   426
  val empty =
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   427
    make_data Symtab.empty
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
   428
      {selectors = Symtab.empty, updates = Symtab.empty,
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
   429
          simpset = HOL_basic_ss, defset = HOL_basic_ss,
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
   430
          foldcong = HOL_basic_ss, unfoldcong = HOL_basic_ss}
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
   431
       Symtab.empty [] Symtab.empty Symtab.empty Symtab.empty Symtab.empty;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16379
diff changeset
   432
  val extend = I;
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33385
diff changeset
   433
  fun merge
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   434
   ({records = recs1,
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   435
     sel_upd =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   436
      {selectors = sels1, updates = upds1,
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   437
       simpset = ss1, defset = ds1,
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   438
       foldcong = fc1, unfoldcong = uc1},
14255
e6e3e3f0deed Records:
schirmer
parents: 14079
diff changeset
   439
     equalities = equalities1,
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
   440
     extinjects = extinjects1,
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
   441
     extsplit = extsplit1,
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   442
     splits = splits1,
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   443
     extfields = extfields1,
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   444
     fieldext = fieldext1},
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   445
    {records = recs2,
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   446
     sel_upd =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   447
      {selectors = sels2, updates = upds2,
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   448
       simpset = ss2, defset = ds2,
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   449
       foldcong = fc2, unfoldcong = uc2},
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
   450
     equalities = equalities2,
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
   451
     extinjects = extinjects2,
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
   452
     extsplit = extsplit2,
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   453
     splits = splits2,
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   454
     extfields = extfields2,
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   455
     fieldext = fieldext2}) =
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   456
    make_data
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   457
      (Symtab.merge (K true) (recs1, recs2))
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   458
      {selectors = Symtab.merge (K true) (sels1, sels2),
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   459
        updates = Symtab.merge (K true) (upds1, upds2),
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
   460
        simpset = Simplifier.merge_ss (ss1, ss2),
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
   461
        defset = Simplifier.merge_ss (ds1, ds2),
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
   462
        foldcong = Simplifier.merge_ss (fc1, fc2),
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
   463
        unfoldcong = Simplifier.merge_ss (uc1, uc2)}
22634
399e4b4835da canonical merge operations
haftmann
parents: 22596
diff changeset
   464
      (Symtab.merge Thm.eq_thm_prop (equalities1, equalities2))
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33385
diff changeset
   465
      (Thm.merge_thms (extinjects1, extinjects2))
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   466
      (Symtab.merge Thm.eq_thm_prop (extsplit1, extsplit2))
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   467
      (Symtab.merge (fn ((a, b, c, d), (w, x, y, z)) =>
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   468
          Thm.eq_thm (a, w) andalso Thm.eq_thm (b, x) andalso
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   469
          Thm.eq_thm (c, y) andalso Thm.eq_thm (d, z)) (splits1, splits2))
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   470
      (Symtab.merge (K true) (extfields1, extfields2))
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   471
      (Symtab.merge (K true) (fieldext1, fieldext2));
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22747
diff changeset
   472
);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   473
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16379
diff changeset
   474
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   475
(* access 'records' *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   476
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   477
val get_info = Symtab.lookup o #records o Data.get;
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   478
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   479
fun the_info thy name =
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   480
  (case get_info thy name of
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   481
    SOME info => info
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   482
  | NONE => error ("Unknown record type " ^ quote name));
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   483
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   484
fun put_record name info =
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   485
  Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   486
    make_data (Symtab.update (name, info) records)
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   487
      sel_upd equalities extinjects extsplit splits extfields fieldext);
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   488
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22747
diff changeset
   489
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   490
(* access 'sel_upd' *)
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   491
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   492
val get_sel_upd = #sel_upd o Data.get;
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   493
17510
5e3ce025e1a5 tuned simprocs;
wenzelm
parents: 17485
diff changeset
   494
val is_selector = Symtab.defined o #selectors o get_sel_upd;
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17377
diff changeset
   495
val get_updates = Symtab.lookup o #updates o get_sel_upd;
35232
f588e1169c8b renamed Simplifier.theory_context to Simplifier.global_context to emphasize that this is not the real thing;
wenzelm
parents: 35149
diff changeset
   496
fun get_ss_with_context getss thy = Simplifier.global_context thy (getss (get_sel_upd thy));
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   497
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   498
val get_simpset = get_ss_with_context #simpset;
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
   499
val get_sel_upd_defs = get_ss_with_context #defset;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
   500
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   501
fun get_update_details u thy =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   502
  let val sel_upd = get_sel_upd thy in
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   503
    (case Symtab.lookup (#updates sel_upd) u of
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   504
      SOME s =>
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   505
        let val SOME (dep, ismore) = Symtab.lookup (#selectors sel_upd) s
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   506
        in SOME (s, dep, ismore) end
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   507
    | NONE => NONE)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   508
  end;
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
   509
32744
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   510
fun put_sel_upd names more depth simps defs (folds, unfolds) thy =
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   511
  let
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   512
    val all = names @ [more];
32744
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   513
    val sels = map (rpair (depth, false)) names @ [(more, (depth, true))];
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   514
    val upds = map (suffix updateN) all ~~ all;
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   515
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   516
    val {records, sel_upd = {selectors, updates, simpset, defset, foldcong, unfoldcong},
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   517
      equalities, extinjects, extsplit, splits, extfields, fieldext} = Data.get thy;
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   518
    val data = make_data records
32744
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   519
      {selectors = fold Symtab.update_new sels selectors,
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   520
        updates = fold Symtab.update_new upds updates,
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   521
        simpset = Simplifier.addsimps (simpset, simps),
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   522
        defset = Simplifier.addsimps (defset, defs),
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   523
        foldcong = foldcong addcongs folds,
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   524
        unfoldcong = unfoldcong addcongs unfolds}
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
   525
       equalities extinjects extsplit splits extfields fieldext;
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   526
  in Data.put data thy end;
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22747
diff changeset
   527
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   528
14079
1c22e5499eeb - record_split_tac now also works for object-level universal quantifier
berghofe
parents: 13904
diff changeset
   529
(* access 'equalities' *)
1c22e5499eeb - record_split_tac now also works for object-level universal quantifier
berghofe
parents: 13904
diff changeset
   530
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   531
fun add_equalities name thm =
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   532
  Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   533
    make_data records sel_upd
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   534
      (Symtab.update_new (name, thm) equalities) extinjects extsplit splits extfields fieldext);
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   535
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   536
val get_equalities = Symtab.lookup o #equalities o Data.get;
14079
1c22e5499eeb - record_split_tac now also works for object-level universal quantifier
berghofe
parents: 13904
diff changeset
   537
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22747
diff changeset
   538
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
   539
(* access 'extinjects' *)
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
   540
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   541
fun add_extinjects thm =
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   542
  Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   543
    make_data records sel_upd equalities (insert Thm.eq_thm_prop thm extinjects)
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   544
      extsplit splits extfields fieldext);
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   545
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   546
val get_extinjects = rev o #extinjects o Data.get;
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
   547
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22747
diff changeset
   548
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
   549
(* access 'extsplit' *)
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
   550
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   551
fun add_extsplit name thm =
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   552
  Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   553
    make_data records sel_upd equalities extinjects
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   554
      (Symtab.update_new (name, thm) extsplit) splits extfields fieldext);
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
   555
26088
9b48d0264ffd added get_parent (for AWE);
wenzelm
parents: 26065
diff changeset
   556
14255
e6e3e3f0deed Records:
schirmer
parents: 14079
diff changeset
   557
(* access 'splits' *)
e6e3e3f0deed Records:
schirmer
parents: 14079
diff changeset
   558
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   559
fun add_splits name thmP =
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   560
  Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   561
    make_data records sel_upd equalities extinjects extsplit
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   562
      (Symtab.update_new (name, thmP) splits) extfields fieldext);
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   563
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   564
val get_splits = Symtab.lookup o #splits o Data.get;
14255
e6e3e3f0deed Records:
schirmer
parents: 14079
diff changeset
   565
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
   566
26088
9b48d0264ffd added get_parent (for AWE);
wenzelm
parents: 26065
diff changeset
   567
(* parent/extension of named record *)
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
   568
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   569
val get_parent = (Option.join o Option.map #parent) oo (Symtab.lookup o #records o Data.get);
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   570
val get_extension = Option.map #extension oo (Symtab.lookup o #records o Data.get);
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
   571
14079
1c22e5499eeb - record_split_tac now also works for object-level universal quantifier
berghofe
parents: 13904
diff changeset
   572
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   573
(* access 'extfields' *)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   574
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   575
fun add_extfields name fields =
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   576
  Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   577
    make_data records sel_upd equalities extinjects extsplit splits
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   578
      (Symtab.update_new (name, fields) extfields) fieldext);
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   579
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   580
val get_extfields = Symtab.lookup o #extfields o Data.get;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   581
18858
ceb93f3af7f0 advanced translations: Context.generic;
wenzelm
parents: 18728
diff changeset
   582
fun get_extT_fields thy T =
15059
schirmer
parents: 15058
diff changeset
   583
  let
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   584
    val ((name, Ts), moreT) = dest_recT T;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   585
    val recname =
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
   586
      let val (nm :: _ :: rst) = rev (Long_Name.explode name)   (* FIXME !? *)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   587
      in Long_Name.implode (rev (nm :: rst)) end;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   588
    val midx = maxidx_of_typs (moreT :: Ts);
19748
5d05d091eecb fixed bug in type print translations
schirmer
parents: 19343
diff changeset
   589
    val varifyT = varifyT midx;
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   590
    val {records, extfields, ...} = Data.get thy;
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   591
    val (fields, (more, _)) = split_last (Symtab.lookup_list extfields name);
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17377
diff changeset
   592
    val args = map varifyT (snd (#extension (the (Symtab.lookup records recname))));
15058
cc8f1de3f86c added: get_extT_fields and
schirmer
parents: 15015
diff changeset
   593
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   594
    val subst = fold (Sign.typ_match thy) (but_last args ~~ but_last Ts) Vartab.empty;
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   595
    val fields' = map (apsnd (Envir.norm_type subst o varifyT)) fields;
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   596
  in (fields', (more, moreT)) end;
15058
cc8f1de3f86c added: get_extT_fields and
schirmer
parents: 15015
diff changeset
   597
18858
ceb93f3af7f0 advanced translations: Context.generic;
wenzelm
parents: 18728
diff changeset
   598
fun get_recT_fields thy T =
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
   599
  let
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   600
    val (root_fields, (root_more, root_moreT)) = get_extT_fields thy T;
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   601
    val (rest_fields, rest_more) =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   602
      if is_recT root_moreT then get_recT_fields thy root_moreT
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   603
      else ([], (root_more, root_moreT));
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   604
  in (root_fields @ rest_fields, rest_more) end;
15059
schirmer
parents: 15058
diff changeset
   605
15058
cc8f1de3f86c added: get_extT_fields and
schirmer
parents: 15015
diff changeset
   606
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   607
(* access 'fieldext' *)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   608
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   609
fun add_fieldext extname_types fields =
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   610
  Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   611
    let
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   612
      val fieldext' =
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   613
        fold (fn field => Symtab.update_new (field, extname_types)) fields fieldext;
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   614
    in make_data records sel_upd equalities extinjects extsplit splits extfields fieldext' end);
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   615
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
   616
val get_fieldext = Symtab.lookup o #fieldext o Data.get;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   617
21962
279b129498b6 removed conditional combinator;
wenzelm
parents: 21858
diff changeset
   618
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   619
(* parent records *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   620
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
   621
fun add_parents _ NONE parents = parents
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15273
diff changeset
   622
  | add_parents thy (SOME (types, name)) parents =
12247
9b029789aff6 derive cases/induct rules for ``more'' parts;
wenzelm
parents: 12129
diff changeset
   623
      let
9b029789aff6 derive cases/induct rules for ``more'' parts;
wenzelm
parents: 12129
diff changeset
   624
        fun err msg = error (msg ^ " parent record " ^ quote name);
12255
wenzelm
parents: 12247
diff changeset
   625
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   626
        val {args, parent, fields, extension, induct_scheme, ext_def, ...} =
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   627
          (case get_info thy name of SOME info => info | NONE => err "Unknown");
12247
9b029789aff6 derive cases/induct rules for ``more'' parts;
wenzelm
parents: 12129
diff changeset
   628
        val _ = if length types <> length args then err "Bad number of arguments for" else ();
12255
wenzelm
parents: 12247
diff changeset
   629
12247
9b029789aff6 derive cases/induct rules for ``more'' parts;
wenzelm
parents: 12129
diff changeset
   630
        fun bad_inst ((x, S), T) =
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 22219
diff changeset
   631
          if Sign.of_sort thy (T, S) then NONE else SOME x
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32809
diff changeset
   632
        val bads = map_filter bad_inst (args ~~ types);
21962
279b129498b6 removed conditional combinator;
wenzelm
parents: 21858
diff changeset
   633
        val _ = null bads orelse err ("Ill-sorted instantiation of " ^ commas bads ^ " in");
12255
wenzelm
parents: 12247
diff changeset
   634
12247
9b029789aff6 derive cases/induct rules for ``more'' parts;
wenzelm
parents: 12129
diff changeset
   635
        val inst = map fst args ~~ types;
17377
afaa031ed4da introduced AList.lookup
haftmann
parents: 17337
diff changeset
   636
        val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst);
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   637
        val parent' = Option.map (apfst (map subst)) parent;
12247
9b029789aff6 derive cases/induct rules for ``more'' parts;
wenzelm
parents: 12129
diff changeset
   638
        val fields' = map (apsnd subst) fields;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   639
        val extension' = apsnd (map subst) extension;
12247
9b029789aff6 derive cases/induct rules for ``more'' parts;
wenzelm
parents: 12129
diff changeset
   640
      in
12255
wenzelm
parents: 12247
diff changeset
   641
        add_parents thy parent'
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
   642
          (make_parent_info name fields' extension' ext_def induct_scheme :: parents)
12247
9b029789aff6 derive cases/induct rules for ``more'' parts;
wenzelm
parents: 12129
diff changeset
   643
      end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   644
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
   645
21962
279b129498b6 removed conditional combinator;
wenzelm
parents: 21858
diff changeset
   646
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   647
(** concrete syntax for records **)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   648
22693
fb5e080fa82b adapted decode_type;
wenzelm
parents: 22675
diff changeset
   649
(* decode type *)
fb5e080fa82b adapted decode_type;
wenzelm
parents: 22675
diff changeset
   650
fb5e080fa82b adapted decode_type;
wenzelm
parents: 22675
diff changeset
   651
fun decode_type thy t =
fb5e080fa82b adapted decode_type;
wenzelm
parents: 22675
diff changeset
   652
  let
35144
wenzelm
parents: 35142
diff changeset
   653
    fun get_sort env xi =
wenzelm
parents: 35142
diff changeset
   654
      the_default (Sign.defaultS thy) (AList.lookup (op =) env (xi: indexname));
22693
fb5e080fa82b adapted decode_type;
wenzelm
parents: 22675
diff changeset
   655
  in
35430
df2862dc23a8 adapted to authentic syntax -- actual types are verbatim;
wenzelm
parents: 35410
diff changeset
   656
    Syntax.typ_of_term (get_sort (Syntax.term_sorts t)) t
22693
fb5e080fa82b adapted decode_type;
wenzelm
parents: 22675
diff changeset
   657
  end;
fb5e080fa82b adapted decode_type;
wenzelm
parents: 22675
diff changeset
   658
fb5e080fa82b adapted decode_type;
wenzelm
parents: 22675
diff changeset
   659
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   660
(* parse translations *)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   661
35144
wenzelm
parents: 35142
diff changeset
   662
local
wenzelm
parents: 35142
diff changeset
   663
35146
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   664
fun field_type_tr ((Const (@{syntax_const "_field_type"}, _) $ Const (name, _) $ arg)) =
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   665
      (name, arg)
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   666
  | field_type_tr t = raise TERM ("field_type_tr", [t]);
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   667
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   668
fun field_types_tr (Const (@{syntax_const "_field_types"}, _) $ t $ u) =
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   669
      field_type_tr t :: field_types_tr u
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   670
  | field_types_tr t = [field_type_tr t];
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   671
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   672
fun record_field_types_tr more ctxt t =
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
   673
  let
21772
7c7ade4f537b advanced translation functions: Proof.context;
wenzelm
parents: 21708
diff changeset
   674
    val thy = ProofContext.theory_of ctxt;
35146
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   675
    fun err msg = raise TERM ("Error in record-type input: " ^ msg, [t]);
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   676
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   677
    fun split_args (field :: fields) ((name, arg) :: fargs) =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   678
          if can (unsuffix name) field then
35146
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   679
            let val (args, rest) = split_args fields fargs
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   680
            in (arg :: args, rest) end
35135
1667fd3b051d tuned errors;
wenzelm
parents: 35133
diff changeset
   681
          else err ("expecting field " ^ field ^ " but got " ^ name)
35146
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   682
      | split_args [] (fargs as (_ :: _)) = ([], fargs)
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   683
      | split_args (_ :: _) [] = err "expecting more fields"
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   684
      | split_args _ _ = ([], []);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   685
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
   686
    fun mk_ext (fargs as (name, _) :: _) =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   687
          (case get_fieldext thy (Sign.intern_const thy name) of
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   688
            SOME (ext, alphas) =>
18858
ceb93f3af7f0 advanced translations: Context.generic;
wenzelm
parents: 18728
diff changeset
   689
              (case get_extfields thy ext of
35146
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   690
                SOME fields =>
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   691
                  let
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   692
                    val fields' = but_last fields;
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   693
                    val types = map snd fields';
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   694
                    val (args, rest) = split_args (map fst fields') fargs;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   695
                    val argtypes = map (Sign.certify_typ thy o decode_type thy) args;
33029
2fefe039edf1 uniform use of Integer.min/max;
wenzelm
parents: 32977
diff changeset
   696
                    val midx = fold Term.maxidx_typ argtypes 0;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   697
                    val varifyT = varifyT midx;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   698
                    val vartypes = map varifyT types;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   699
36159
bffb04bf4e83 more robust record syntax: use Type.raw_match to ignore sort constraints as in regular abbreviations (also note that constraints only affect operations, not types);
wenzelm
parents: 36153
diff changeset
   700
                    val subst = Type.raw_matches (vartypes, argtypes) Vartab.empty
35146
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   701
                      handle Type.TYPE_MATCH => err "type is no proper record (extension)";
39134
917b4b6ba3d2 turned show_sorts/show_types into proper configuration options;
wenzelm
parents: 38864
diff changeset
   702
                    val term_of_typ = Syntax.term_of_typ (Config.get ctxt show_sorts);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   703
                    val alphas' =
39134
917b4b6ba3d2 turned show_sorts/show_types into proper configuration options;
wenzelm
parents: 38864
diff changeset
   704
                      map (term_of_typ o Envir.norm_type subst o varifyT) (but_last alphas);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   705
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   706
                    val more' = mk_ext rest;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   707
                  in
35430
df2862dc23a8 adapted to authentic syntax -- actual types are verbatim;
wenzelm
parents: 35410
diff changeset
   708
                    list_comb
df2862dc23a8 adapted to authentic syntax -- actual types are verbatim;
wenzelm
parents: 35410
diff changeset
   709
                      (Syntax.const (Syntax.mark_type (suffix ext_typeN ext)), alphas' @ [more'])
35146
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   710
                  end
35135
1667fd3b051d tuned errors;
wenzelm
parents: 35133
diff changeset
   711
              | NONE => err ("no fields defined for " ^ ext))
1667fd3b051d tuned errors;
wenzelm
parents: 35133
diff changeset
   712
          | NONE => err (name ^ " is no proper field"))
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   713
      | mk_ext [] = more;
35146
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   714
  in
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   715
    mk_ext (field_types_tr t)
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   716
  end;
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   717
35363
09489d8ffece explicit @{type_syntax} markup;
wenzelm
parents: 35262
diff changeset
   718
fun record_type_tr ctxt [t] = record_field_types_tr (Syntax.const @{type_syntax unit}) ctxt t
35146
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   719
  | record_type_tr _ ts = raise TERM ("record_type_tr", ts);
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   720
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   721
fun record_type_scheme_tr ctxt [t, more] = record_field_types_tr more ctxt t
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   722
  | record_type_scheme_tr _ ts = raise TERM ("record_type_scheme_tr", ts);
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   723
35147
wenzelm
parents: 35146
diff changeset
   724
wenzelm
parents: 35146
diff changeset
   725
fun field_tr ((Const (@{syntax_const "_field"}, _) $ Const (name, _) $ arg)) = (name, arg)
wenzelm
parents: 35146
diff changeset
   726
  | field_tr t = raise TERM ("field_tr", [t]);
wenzelm
parents: 35146
diff changeset
   727
wenzelm
parents: 35146
diff changeset
   728
fun fields_tr (Const (@{syntax_const "_fields"}, _) $ t $ u) = field_tr t :: fields_tr u
wenzelm
parents: 35146
diff changeset
   729
  | fields_tr t = [field_tr t];
wenzelm
parents: 35146
diff changeset
   730
wenzelm
parents: 35146
diff changeset
   731
fun record_fields_tr more ctxt t =
wenzelm
parents: 35146
diff changeset
   732
  let
wenzelm
parents: 35146
diff changeset
   733
    val thy = ProofContext.theory_of ctxt;
wenzelm
parents: 35146
diff changeset
   734
    fun err msg = raise TERM ("Error in record input: " ^ msg, [t]);
wenzelm
parents: 35146
diff changeset
   735
wenzelm
parents: 35146
diff changeset
   736
    fun split_args (field :: fields) ((name, arg) :: fargs) =
wenzelm
parents: 35146
diff changeset
   737
          if can (unsuffix name) field
wenzelm
parents: 35146
diff changeset
   738
          then
wenzelm
parents: 35146
diff changeset
   739
            let val (args, rest) = split_args fields fargs
wenzelm
parents: 35146
diff changeset
   740
            in (arg :: args, rest) end
wenzelm
parents: 35146
diff changeset
   741
          else err ("expecting field " ^ field ^ " but got " ^ name)
wenzelm
parents: 35146
diff changeset
   742
      | split_args [] (fargs as (_ :: _)) = ([], fargs)
wenzelm
parents: 35146
diff changeset
   743
      | split_args (_ :: _) [] = err "expecting more fields"
wenzelm
parents: 35146
diff changeset
   744
      | split_args _ _ = ([], []);
wenzelm
parents: 35146
diff changeset
   745
wenzelm
parents: 35146
diff changeset
   746
    fun mk_ext (fargs as (name, _) :: _) =
wenzelm
parents: 35146
diff changeset
   747
          (case get_fieldext thy (Sign.intern_const thy name) of
wenzelm
parents: 35146
diff changeset
   748
            SOME (ext, _) =>
wenzelm
parents: 35146
diff changeset
   749
              (case get_extfields thy ext of
wenzelm
parents: 35146
diff changeset
   750
                SOME fields =>
wenzelm
parents: 35146
diff changeset
   751
                  let
wenzelm
parents: 35146
diff changeset
   752
                    val (args, rest) = split_args (map fst (but_last fields)) fargs;
wenzelm
parents: 35146
diff changeset
   753
                    val more' = mk_ext rest;
35262
9ea4445d2ccf slightly more abstract syntax mark/unmark operations;
wenzelm
parents: 35240
diff changeset
   754
                  in list_comb (Syntax.const (Syntax.mark_const (ext ^ extN)), args @ [more']) end
35147
wenzelm
parents: 35146
diff changeset
   755
              | NONE => err ("no fields defined for " ^ ext))
wenzelm
parents: 35146
diff changeset
   756
          | NONE => err (name ^ " is no proper field"))
wenzelm
parents: 35146
diff changeset
   757
      | mk_ext [] = more;
wenzelm
parents: 35146
diff changeset
   758
  in mk_ext (fields_tr t) end;
wenzelm
parents: 35146
diff changeset
   759
wenzelm
parents: 35146
diff changeset
   760
fun record_tr ctxt [t] = record_fields_tr (Syntax.const @{const_syntax Unity}) ctxt t
wenzelm
parents: 35146
diff changeset
   761
  | record_tr _ ts = raise TERM ("record_tr", ts);
wenzelm
parents: 35146
diff changeset
   762
wenzelm
parents: 35146
diff changeset
   763
fun record_scheme_tr ctxt [t, more] = record_fields_tr more ctxt t
wenzelm
parents: 35146
diff changeset
   764
  | record_scheme_tr _ ts = raise TERM ("record_scheme_tr", ts);
wenzelm
parents: 35146
diff changeset
   765
wenzelm
parents: 35146
diff changeset
   766
wenzelm
parents: 35146
diff changeset
   767
fun field_update_tr (Const (@{syntax_const "_field_update"}, _) $ Const (name, _) $ arg) =
wenzelm
parents: 35146
diff changeset
   768
      Syntax.const (suffix updateN name) $ Abs ("_", dummyT, arg)
wenzelm
parents: 35146
diff changeset
   769
  | field_update_tr t = raise TERM ("field_update_tr", [t]);
wenzelm
parents: 35146
diff changeset
   770
wenzelm
parents: 35146
diff changeset
   771
fun field_updates_tr (Const (@{syntax_const "_field_updates"}, _) $ t $ u) =
wenzelm
parents: 35146
diff changeset
   772
      field_update_tr t :: field_updates_tr u
wenzelm
parents: 35146
diff changeset
   773
  | field_updates_tr t = [field_update_tr t];
wenzelm
parents: 35146
diff changeset
   774
wenzelm
parents: 35146
diff changeset
   775
fun record_update_tr [t, u] = fold (curry op $) (field_updates_tr u) t
wenzelm
parents: 35146
diff changeset
   776
  | record_update_tr ts = raise TERM ("record_update_tr", ts);
wenzelm
parents: 35146
diff changeset
   777
35146
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   778
in
15215
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
   779
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24830
diff changeset
   780
val parse_translation =
35145
f132a4fd8679 moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents: 35144
diff changeset
   781
 [(@{syntax_const "_record_update"}, record_update_tr)];
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   782
35146
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   783
val advanced_parse_translation =
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   784
 [(@{syntax_const "_record"}, record_tr),
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   785
  (@{syntax_const "_record_scheme"}, record_scheme_tr),
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   786
  (@{syntax_const "_record_type"}, record_type_tr),
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   787
  (@{syntax_const "_record_type_scheme"}, record_type_scheme_tr)];
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   788
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
   789
end;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   790
21226
a607ae87ee81 field-update in records is generalised to take a function on the field
schirmer
parents: 21109
diff changeset
   791
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   792
(* print translations *)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   793
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   794
val print_type_abbr = Unsynchronized.ref true;
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   795
val print_type_as_fields = Unsynchronized.ref true;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   796
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   797
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   798
local
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   799
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   800
(* FIXME early extern (!??) *)
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   801
(* FIXME Syntax.free (??) *)
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   802
fun field_type_tr' (c, t) = Syntax.const @{syntax_const "_field_type"} $ Syntax.const c $ t;
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   803
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   804
fun field_types_tr' (t, u) = Syntax.const @{syntax_const "_field_types"} $ t $ u;
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   805
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   806
fun record_type_tr' ctxt t =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   807
  let
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   808
    val thy = ProofContext.theory_of ctxt;
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   809
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   810
    val T = decode_type thy t;
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   811
    val varifyT = varifyT (Term.maxidx_of_typ T);
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   812
39134
917b4b6ba3d2 turned show_sorts/show_types into proper configuration options;
wenzelm
parents: 38864
diff changeset
   813
    val term_of_type = Syntax.term_of_typ (Config.get ctxt show_sorts);
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   814
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   815
    fun strip_fields T =
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   816
      (case T of
35615
61bb9f8af129 record_type_tr': more robust strip_fields (printed types are not necessarily well-formed, e.g. in Syntax.pretty_arity);
wenzelm
parents: 35614
diff changeset
   817
        Type (ext, args as _ :: _) =>
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   818
          (case try (unsuffix ext_typeN) ext of
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   819
            SOME ext' =>
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   820
              (case get_extfields thy ext' of
35615
61bb9f8af129 record_type_tr': more robust strip_fields (printed types are not necessarily well-formed, e.g. in Syntax.pretty_arity);
wenzelm
parents: 35614
diff changeset
   821
                SOME (fields as (x, _) :: _) =>
61bb9f8af129 record_type_tr': more robust strip_fields (printed types are not necessarily well-formed, e.g. in Syntax.pretty_arity);
wenzelm
parents: 35614
diff changeset
   822
                  (case get_fieldext thy x of
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   823
                    SOME (_, alphas) =>
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   824
                     (let
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   825
                        val f :: fs = but_last fields;
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   826
                        val fields' =
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   827
                          apfst (Sign.extern_const thy) f :: map (apfst Long_Name.base_name) fs;
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   828
                        val (args', more) = split_last args;
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   829
                        val alphavars = map varifyT (but_last alphas);
36159
bffb04bf4e83 more robust record syntax: use Type.raw_match to ignore sort constraints as in regular abbreviations (also note that constraints only affect operations, not types);
wenzelm
parents: 36153
diff changeset
   830
                        val subst = Type.raw_matches (alphavars, args') Vartab.empty;
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   831
                        val fields'' = (map o apsnd) (Envir.norm_type subst o varifyT) fields';
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   832
                      in fields'' @ strip_fields more end
36159
bffb04bf4e83 more robust record syntax: use Type.raw_match to ignore sort constraints as in regular abbreviations (also note that constraints only affect operations, not types);
wenzelm
parents: 36153
diff changeset
   833
                      handle Type.TYPE_MATCH => [("", T)])
35615
61bb9f8af129 record_type_tr': more robust strip_fields (printed types are not necessarily well-formed, e.g. in Syntax.pretty_arity);
wenzelm
parents: 35614
diff changeset
   834
                  | _ => [("", T)])
61bb9f8af129 record_type_tr': more robust strip_fields (printed types are not necessarily well-formed, e.g. in Syntax.pretty_arity);
wenzelm
parents: 35614
diff changeset
   835
              | _ => [("", T)])
61bb9f8af129 record_type_tr': more robust strip_fields (printed types are not necessarily well-formed, e.g. in Syntax.pretty_arity);
wenzelm
parents: 35614
diff changeset
   836
          | _ => [("", T)])
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   837
      | _ => [("", T)]);
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   838
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   839
    val (fields, (_, moreT)) = split_last (strip_fields T);
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   840
    val _ = null fields andalso raise Match;
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   841
    val u = foldr1 field_types_tr' (map (field_type_tr' o apsnd term_of_type) fields);
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   842
  in
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   843
    if not (! print_type_as_fields) orelse null fields then raise Match
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   844
    else if moreT = HOLogic.unitT then Syntax.const @{syntax_const "_record_type"} $ u
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   845
    else Syntax.const @{syntax_const "_record_type_scheme"} $ u $ term_of_type moreT
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   846
  end;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   847
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   848
(*try to reconstruct the record name type abbreviation from
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   849
  the (nested) extension types*)
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   850
fun record_type_abbr_tr' abbr alphas zeta last_ext schemeT ctxt tm =
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   851
  let
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   852
    val thy = ProofContext.theory_of ctxt;
35614
d7afa8700622 record_type_abbr_tr': removed obsolete workaround for decode_type, which now retains syntactic categories of variables vs. constructors (authentic syntax);
wenzelm
parents: 35430
diff changeset
   853
    val T = decode_type thy tm;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   854
    val midx = maxidx_of_typ T;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   855
    val varifyT = varifyT midx;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   856
36151
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
   857
    fun mk_type_abbr subst name args =
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
   858
      let val abbrT = Type (name, map (varifyT o TFree) args)
39134
917b4b6ba3d2 turned show_sorts/show_types into proper configuration options;
wenzelm
parents: 38864
diff changeset
   859
      in Syntax.term_of_typ (Config.get ctxt show_sorts) (Envir.norm_type subst abbrT) end;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   860
36159
bffb04bf4e83 more robust record syntax: use Type.raw_match to ignore sort constraints as in regular abbreviations (also note that constraints only affect operations, not types);
wenzelm
parents: 36153
diff changeset
   861
    fun match rT T = Type.raw_match (varifyT rT, T) Vartab.empty;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   862
  in
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
   863
    if ! print_type_abbr then
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   864
      (case last_extT T of
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   865
        SOME (name, _) =>
35148
3a34fee2cfcd eliminated camel case;
wenzelm
parents: 35147
diff changeset
   866
          if name = last_ext then
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   867
            let val subst = match schemeT T in
36151
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
   868
              if HOLogic.is_unitT (Envir.norm_type subst (varifyT (TFree zeta)))
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
   869
              then mk_type_abbr subst abbr alphas
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
   870
              else mk_type_abbr subst (suffix schemeN abbr) (alphas @ [zeta])
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   871
            end handle Type.TYPE_MATCH => record_type_tr' ctxt tm
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   872
          else raise Match (*give print translation of specialised record a chance*)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   873
      | _ => raise Match)
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   874
    else record_type_tr' ctxt tm
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   875
  end;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   876
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   877
in
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   878
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   879
fun record_ext_type_tr' name =
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   880
  let
35430
df2862dc23a8 adapted to authentic syntax -- actual types are verbatim;
wenzelm
parents: 35410
diff changeset
   881
    val ext_type_name = Syntax.mark_type (suffix ext_typeN name);
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   882
    fun tr' ctxt ts =
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   883
      record_type_tr' ctxt (list_comb (Syntax.const ext_type_name, ts));
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   884
  in (ext_type_name, tr') end;
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   885
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   886
fun record_ext_type_abbr_tr' abbr alphas zeta last_ext schemeT name =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   887
  let
35430
df2862dc23a8 adapted to authentic syntax -- actual types are verbatim;
wenzelm
parents: 35410
diff changeset
   888
    val ext_type_name = Syntax.mark_type (suffix ext_typeN name);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
   889
    fun tr' ctxt ts =
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   890
      record_type_abbr_tr' abbr alphas zeta last_ext schemeT ctxt
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   891
        (list_comb (Syntax.const ext_type_name, ts));
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   892
  in (ext_type_name, tr') end;
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   893
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
   894
end;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
   895
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
   896
35240
wenzelm
parents: 35239
diff changeset
   897
local
wenzelm
parents: 35239
diff changeset
   898
wenzelm
parents: 35239
diff changeset
   899
(* FIXME Syntax.free (??) *)
wenzelm
parents: 35239
diff changeset
   900
fun field_tr' (c, t) = Syntax.const @{syntax_const "_field"} $ Syntax.const c $ t;
wenzelm
parents: 35239
diff changeset
   901
fun fields_tr' (t, u) = Syntax.const @{syntax_const "_fields"} $ t $ u;
wenzelm
parents: 35239
diff changeset
   902
wenzelm
parents: 35239
diff changeset
   903
fun record_tr' ctxt t =
wenzelm
parents: 35239
diff changeset
   904
  let
wenzelm
parents: 35239
diff changeset
   905
    val thy = ProofContext.theory_of ctxt;
wenzelm
parents: 35239
diff changeset
   906
    val extern = Consts.extern (ProofContext.consts_of ctxt);
wenzelm
parents: 35239
diff changeset
   907
wenzelm
parents: 35239
diff changeset
   908
    fun strip_fields t =
wenzelm
parents: 35239
diff changeset
   909
      (case strip_comb t of
wenzelm
parents: 35239
diff changeset
   910
        (Const (ext, _), args as (_ :: _)) =>
35262
9ea4445d2ccf slightly more abstract syntax mark/unmark operations;
wenzelm
parents: 35240
diff changeset
   911
          (case try (Syntax.unmark_const o unsuffix extN) ext of
35240
wenzelm
parents: 35239
diff changeset
   912
            SOME ext' =>
wenzelm
parents: 35239
diff changeset
   913
              (case get_extfields thy ext' of
wenzelm
parents: 35239
diff changeset
   914
                SOME fields =>
wenzelm
parents: 35239
diff changeset
   915
                 (let
wenzelm
parents: 35239
diff changeset
   916
                    val f :: fs = but_last (map fst fields);
wenzelm
parents: 35239
diff changeset
   917
                    val fields' = extern f :: map Long_Name.base_name fs;
wenzelm
parents: 35239
diff changeset
   918
                    val (args', more) = split_last args;
wenzelm
parents: 35239
diff changeset
   919
                  in (fields' ~~ args') @ strip_fields more end
wenzelm
parents: 35239
diff changeset
   920
                  handle Library.UnequalLengths => [("", t)])
wenzelm
parents: 35239
diff changeset
   921
              | NONE => [("", t)])
wenzelm
parents: 35239
diff changeset
   922
          | NONE => [("", t)])
wenzelm
parents: 35239
diff changeset
   923
       | _ => [("", t)]);
wenzelm
parents: 35239
diff changeset
   924
wenzelm
parents: 35239
diff changeset
   925
    val (fields, (_, more)) = split_last (strip_fields t);
wenzelm
parents: 35239
diff changeset
   926
    val _ = null fields andalso raise Match;
wenzelm
parents: 35239
diff changeset
   927
    val u = foldr1 fields_tr' (map field_tr' fields);
wenzelm
parents: 35239
diff changeset
   928
  in
wenzelm
parents: 35239
diff changeset
   929
    case more of
wenzelm
parents: 35239
diff changeset
   930
      Const (@{const_syntax Unity}, _) => Syntax.const @{syntax_const "_record"} $ u
wenzelm
parents: 35239
diff changeset
   931
    | _ => Syntax.const @{syntax_const "_record_scheme"} $ u $ more
wenzelm
parents: 35239
diff changeset
   932
  end;
wenzelm
parents: 35239
diff changeset
   933
wenzelm
parents: 35239
diff changeset
   934
in
wenzelm
parents: 35239
diff changeset
   935
wenzelm
parents: 35239
diff changeset
   936
fun record_ext_tr' name =
wenzelm
parents: 35239
diff changeset
   937
  let
35262
9ea4445d2ccf slightly more abstract syntax mark/unmark operations;
wenzelm
parents: 35240
diff changeset
   938
    val ext_name = Syntax.mark_const (name ^ extN);
35240
wenzelm
parents: 35239
diff changeset
   939
    fun tr' ctxt ts = record_tr' ctxt (list_comb (Syntax.const ext_name, ts));
wenzelm
parents: 35239
diff changeset
   940
  in (ext_name, tr') end;
wenzelm
parents: 35239
diff changeset
   941
wenzelm
parents: 35239
diff changeset
   942
end;
wenzelm
parents: 35239
diff changeset
   943
wenzelm
parents: 35239
diff changeset
   944
wenzelm
parents: 35239
diff changeset
   945
local
wenzelm
parents: 35239
diff changeset
   946
wenzelm
parents: 35239
diff changeset
   947
fun field_updates_tr' ctxt (tm as Const (c, _) $ k $ u) =
wenzelm
parents: 35239
diff changeset
   948
      let
wenzelm
parents: 35239
diff changeset
   949
        val extern = Consts.extern (ProofContext.consts_of ctxt);
wenzelm
parents: 35239
diff changeset
   950
        val t =
wenzelm
parents: 35239
diff changeset
   951
          (case k of
wenzelm
parents: 35239
diff changeset
   952
            Abs (_, _, Abs (_, _, t) $ Bound 0) =>
wenzelm
parents: 35239
diff changeset
   953
              if null (loose_bnos t) then t else raise Match
wenzelm
parents: 35239
diff changeset
   954
          | Abs (_, _, t) =>
wenzelm
parents: 35239
diff changeset
   955
              if null (loose_bnos t) then t else raise Match
wenzelm
parents: 35239
diff changeset
   956
          | _ => raise Match);
wenzelm
parents: 35239
diff changeset
   957
      in
35262
9ea4445d2ccf slightly more abstract syntax mark/unmark operations;
wenzelm
parents: 35240
diff changeset
   958
        (case Option.map extern (try Syntax.unmark_const c) of
35240
wenzelm
parents: 35239
diff changeset
   959
          SOME update_name =>
wenzelm
parents: 35239
diff changeset
   960
            (case try (unsuffix updateN) update_name of
wenzelm
parents: 35239
diff changeset
   961
              SOME name =>
wenzelm
parents: 35239
diff changeset
   962
                apfst (cons (Syntax.const @{syntax_const "_field_update"} $ Syntax.free name $ t))
wenzelm
parents: 35239
diff changeset
   963
                  (field_updates_tr' ctxt u)
wenzelm
parents: 35239
diff changeset
   964
            | NONE => ([], tm))
wenzelm
parents: 35239
diff changeset
   965
        | NONE => ([], tm))
wenzelm
parents: 35239
diff changeset
   966
      end
wenzelm
parents: 35239
diff changeset
   967
  | field_updates_tr' _ tm = ([], tm);
wenzelm
parents: 35239
diff changeset
   968
wenzelm
parents: 35239
diff changeset
   969
fun record_update_tr' ctxt tm =
wenzelm
parents: 35239
diff changeset
   970
  (case field_updates_tr' ctxt tm of
wenzelm
parents: 35239
diff changeset
   971
    ([], _) => raise Match
wenzelm
parents: 35239
diff changeset
   972
  | (ts, u) =>
wenzelm
parents: 35239
diff changeset
   973
      Syntax.const @{syntax_const "_record_update"} $ u $
wenzelm
parents: 35239
diff changeset
   974
        foldr1 (fn (v, w) => Syntax.const @{syntax_const "_field_updates"} $ v $ w) (rev ts));
wenzelm
parents: 35239
diff changeset
   975
wenzelm
parents: 35239
diff changeset
   976
in
wenzelm
parents: 35239
diff changeset
   977
wenzelm
parents: 35239
diff changeset
   978
fun field_update_tr' name =
wenzelm
parents: 35239
diff changeset
   979
  let
35262
9ea4445d2ccf slightly more abstract syntax mark/unmark operations;
wenzelm
parents: 35240
diff changeset
   980
    val update_name = Syntax.mark_const (name ^ updateN);
35240
wenzelm
parents: 35239
diff changeset
   981
    fun tr' ctxt [t, u] = record_update_tr' ctxt (Syntax.const update_name $ t $ u)
wenzelm
parents: 35239
diff changeset
   982
      | tr' _ _ = raise Match;
wenzelm
parents: 35239
diff changeset
   983
  in (update_name, tr') end;
wenzelm
parents: 35239
diff changeset
   984
wenzelm
parents: 35239
diff changeset
   985
end;
wenzelm
parents: 35239
diff changeset
   986
wenzelm
parents: 35239
diff changeset
   987
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
   988
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
   989
(** record simprocs **)
14358
233c5bd5b539 cleaning up
schirmer
parents: 14255
diff changeset
   990
33691
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
   991
fun future_forward_prf_standard thy prf prop () =
33711
2fdb11580b96 guard future proofs by Goal.future_enabled;
wenzelm
parents: 33691
diff changeset
   992
  let val thm =
2fdb11580b96 guard future proofs by Goal.future_enabled;
wenzelm
parents: 33691
diff changeset
   993
    if ! quick_and_dirty then Skip_Proof.make_thm thy prop
2fdb11580b96 guard future proofs by Goal.future_enabled;
wenzelm
parents: 33691
diff changeset
   994
    else if Goal.future_enabled () then
37186
349e9223c685 explicit markup for forked goals, as indicated by Goal.fork;
wenzelm
parents: 37136
diff changeset
   995
      Goal.future_result (ProofContext.init_global thy) (Goal.fork prf) prop
33711
2fdb11580b96 guard future proofs by Goal.future_enabled;
wenzelm
parents: 33691
diff changeset
   996
    else prf ()
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 34151
diff changeset
   997
  in Drule.export_without_context thm end;
33691
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
   998
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
   999
fun prove_common immediate stndrd thy asms prop tac =
33711
2fdb11580b96 guard future proofs by Goal.future_enabled;
wenzelm
parents: 33691
diff changeset
  1000
  let
2fdb11580b96 guard future proofs by Goal.future_enabled;
wenzelm
parents: 33691
diff changeset
  1001
    val prv =
2fdb11580b96 guard future proofs by Goal.future_enabled;
wenzelm
parents: 33691
diff changeset
  1002
      if ! quick_and_dirty then Skip_Proof.prove
2fdb11580b96 guard future proofs by Goal.future_enabled;
wenzelm
parents: 33691
diff changeset
  1003
      else if immediate orelse not (Goal.future_enabled ()) then Goal.prove
2fdb11580b96 guard future proofs by Goal.future_enabled;
wenzelm
parents: 33691
diff changeset
  1004
      else Goal.prove_future;
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36173
diff changeset
  1005
    val prf = prv (ProofContext.init_global thy) [] asms prop tac;
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 34151
diff changeset
  1006
  in if stndrd then Drule.export_without_context prf else prf end;
33691
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  1007
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  1008
val prove_future_global = prove_common false;
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  1009
val prove_global = prove_common true;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1010
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1011
fun is_sel_upd_pair thy (Const (s, _)) (Const (u, t')) =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1012
  (case get_updates thy u of
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1013
    SOME u_name => u_name = s
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1014
  | NONE => raise TERM ("is_sel_upd_pair: not update", [Const (u, t')]));
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1015
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1016
fun mk_comp f g =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1017
  let
32974
2a1aaa2d9e64 eliminated old List.foldr and OldTerm operations;
wenzelm
parents: 32973
diff changeset
  1018
    val X = fastype_of g;
2a1aaa2d9e64 eliminated old List.foldr and OldTerm operations;
wenzelm
parents: 32973
diff changeset
  1019
    val A = domain_type X;
2a1aaa2d9e64 eliminated old List.foldr and OldTerm operations;
wenzelm
parents: 32973
diff changeset
  1020
    val B = range_type X;
2a1aaa2d9e64 eliminated old List.foldr and OldTerm operations;
wenzelm
parents: 32973
diff changeset
  1021
    val C = range_type (fastype_of f);
2a1aaa2d9e64 eliminated old List.foldr and OldTerm operations;
wenzelm
parents: 32973
diff changeset
  1022
    val T = (B --> C) --> (A --> B) --> A --> C;
35133
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1023
  in Const (@{const_name Fun.comp}, T) $ f $ g end;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1024
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1025
fun mk_comp_id f =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1026
  let val T = range_type (fastype_of f)
35133
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1027
  in mk_comp (Const (@{const_name Fun.id}, T --> T)) f end;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1028
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
  1029
fun get_upd_funs (upd $ _ $ t) = upd :: get_upd_funs t
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1030
  | get_upd_funs _ = [];
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1031
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1032
fun get_accupd_simps thy term defset =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1033
  let
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1034
    val (acc, [body]) = strip_comb term;
35408
b48ab741683b modernized structure Term_Ord;
wenzelm
parents: 35363
diff changeset
  1035
    val upd_funs = sort_distinct Term_Ord.fast_term_ord (get_upd_funs body);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1036
    fun get_simp upd =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1037
      let
35133
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1038
        (* FIXME fresh "f" (!?) *)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1039
        val T = domain_type (fastype_of upd);
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1040
        val lhs = mk_comp acc (upd $ Free ("f", T));
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1041
        val rhs =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1042
          if is_sel_upd_pair thy acc upd
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1043
          then mk_comp (Free ("f", T)) acc
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1044
          else mk_comp_id acc;
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1045
        val prop = lhs === rhs;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1046
        val othm =
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36173
diff changeset
  1047
          Goal.prove (ProofContext.init_global thy) [] [] prop
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1048
            (fn _ =>
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1049
              simp_tac defset 1 THEN
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  1050
              REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac 1) THEN
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1051
              TRY (simp_tac (HOL_ss addsimps id_o_apps) 1));
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1052
        val dest =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1053
          if is_sel_upd_pair thy acc upd
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1054
          then o_eq_dest
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1055
          else o_eq_id_dest;
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 34151
diff changeset
  1056
      in Drule.export_without_context (othm RS dest) end;
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
  1057
  in map get_simp upd_funs end;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1058
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1059
fun get_updupd_simp thy defset u u' comp =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1060
  let
35133
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1061
    (* FIXME fresh "f" (!?) *)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1062
    val f = Free ("f", domain_type (fastype_of u));
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1063
    val f' = Free ("f'", domain_type (fastype_of u'));
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1064
    val lhs = mk_comp (u $ f) (u' $ f');
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1065
    val rhs =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1066
      if comp
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1067
      then u $ mk_comp f f'
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1068
      else mk_comp (u' $ f') (u $ f);
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1069
    val prop = lhs === rhs;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1070
    val othm =
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36173
diff changeset
  1071
      Goal.prove (ProofContext.init_global thy) [] [] prop
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1072
        (fn _ =>
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1073
          simp_tac defset 1 THEN
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  1074
          REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac 1) THEN
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1075
          TRY (simp_tac (HOL_ss addsimps [id_apply]) 1));
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1076
    val dest = if comp then o_eq_dest_lhs else o_eq_dest;
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 34151
diff changeset
  1077
  in Drule.export_without_context (othm RS dest) end;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1078
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1079
fun get_updupd_simps thy term defset =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1080
  let
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1081
    val upd_funs = get_upd_funs term;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1082
    val cname = fst o dest_Const;
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1083
    fun getswap u u' = get_updupd_simp thy defset u u' (cname u = cname u');
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1084
    fun build_swaps_to_eq _ [] swaps = swaps
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1085
      | build_swaps_to_eq upd (u :: us) swaps =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1086
          let
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1087
            val key = (cname u, cname upd);
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1088
            val newswaps =
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
  1089
              if Symreltab.defined swaps key then swaps
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
  1090
              else Symreltab.insert (K true) (key, getswap u upd) swaps;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1091
          in
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1092
            if cname u = cname upd then newswaps
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1093
            else build_swaps_to_eq upd us newswaps
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1094
          end;
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1095
    fun swaps_needed [] _ _ swaps = map snd (Symreltab.dest swaps)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1096
      | swaps_needed (u :: us) prev seen swaps =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1097
          if Symtab.defined seen (cname u)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1098
          then swaps_needed us prev seen (build_swaps_to_eq u prev swaps)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1099
          else swaps_needed us (u :: prev) (Symtab.insert (K true) (cname u, ()) seen) swaps;
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
  1100
  in swaps_needed upd_funs [] Symtab.empty Symreltab.empty end;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1101
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  1102
val named_cterm_instantiate = Iso_Tuple_Support.named_cterm_instantiate;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1103
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1104
fun prove_unfold_defs thy ex_simps ex_simprs prop =
21226
a607ae87ee81 field-update in records is generalised to take a function on the field
schirmer
parents: 21109
diff changeset
  1105
  let
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1106
    val defset = get_sel_upd_defs thy;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1107
    val prop' = Envir.beta_eta_contract prop;
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1108
    val (lhs, _) = Logic.dest_equals (Logic.strip_assums_concl prop');
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1109
    val (_, args) = strip_comb lhs;
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1110
    val simps = (if length args = 1 then get_accupd_simps else get_updupd_simps) thy lhs defset;
16973
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 16934
diff changeset
  1111
  in
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36173
diff changeset
  1112
    Goal.prove (ProofContext.init_global thy) [] [] prop'
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1113
      (fn _ =>
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1114
        simp_tac (HOL_basic_ss addsimps (simps @ [K_record_comp])) 1 THEN
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1115
        TRY (simp_tac (HOL_basic_ss addsimps ex_simps addsimprocs ex_simprs) 1))
15203
4481ada46cbb bug-fix
schirmer
parents: 15059
diff changeset
  1116
  end;
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
  1117
15215
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1118
15059
schirmer
parents: 15058
diff changeset
  1119
local
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1120
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1121
fun eq (s1: string) (s2: string) = (s1 = s2);
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1122
16822
7fa91e6176ed avoiding even more garbage
schirmer
parents: 16783
diff changeset
  1123
fun has_field extfields f T =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1124
  exists (fn (eN, _) => exists (eq f o fst) (Symtab.lookup_list extfields eN)) (dest_recTs T);
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1125
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1126
fun K_skeleton n (T as Type (_, [_, kT])) (b as Bound i) (Abs (x, xT, t)) =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1127
      if null (loose_bnos t) then ((n, kT), (Abs (x, xT, Bound (i + 1)))) else ((n, T), b)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1128
  | K_skeleton n T b _ = ((n, T), b);
25705
45a2ffc5911e replaced K_record by lambda term %x. c
schirmer
parents: 25179
diff changeset
  1129
15059
schirmer
parents: 15058
diff changeset
  1130
in
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1131
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1132
(* simproc *)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1133
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1134
(*
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1135
  Simplify selections of an record update:
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1136
    (1)  S (S_update k r) = k (S r)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1137
    (2)  S (X_update k r) = S r
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1138
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1139
  The simproc skips multiple updates at once, eg:
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1140
   S (X_update x (Y_update y (S_update k r))) = k (S r)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1141
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1142
  But be careful in (2) because of the extensibility of records.
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1143
  - If S is a more-selector we have to make sure that the update on component
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1144
    X does not affect the selected subrecord.
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1145
  - If X is a more-selector we have to make sure that S is not in the updated
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1146
    subrecord.
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1147
*)
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1148
val simproc =
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 38544
diff changeset
  1149
  Simplifier.simproc_global @{theory HOL} "record_simp" ["x"]
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1150
    (fn thy => fn _ => fn t =>
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1151
      (case t of
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1152
        (sel as Const (s, Type (_, [_, rangeS]))) $
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1153
            ((upd as Const (u, Type (_, [_, Type (_, [rT, _])]))) $ k $ r) =>
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1154
          if is_selector thy s andalso is_some (get_updates thy u) then
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1155
            let
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1156
              val {sel_upd = {updates, ...}, extfields, ...} = Data.get thy;
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1157
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1158
              fun mk_eq_terms ((upd as Const (u, Type(_, [kT, _]))) $ k $ r) =
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1159
                    (case Symtab.lookup updates u of
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1160
                      NONE => NONE
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1161
                    | SOME u_name =>
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1162
                        if u_name = s then
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1163
                          (case mk_eq_terms r of
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1164
                            NONE =>
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1165
                              let
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1166
                                val rv = ("r", rT);
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1167
                                val rb = Bound 0;
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1168
                                val (kv, kb) = K_skeleton "k" kT (Bound 1) k;
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1169
                              in SOME (upd $ kb $ rb, kb $ (sel $ rb), [kv, rv]) end
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1170
                          | SOME (trm, trm', vars) =>
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1171
                              let
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1172
                                val (kv, kb) = K_skeleton "k" kT (Bound (length vars)) k;
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1173
                              in SOME (upd $ kb $ trm, kb $ trm', kv :: vars) end)
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1174
                        else if has_field extfields u_name rangeS orelse
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1175
                          has_field extfields s (domain_type kT) then NONE
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1176
                        else
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1177
                          (case mk_eq_terms r of
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1178
                            SOME (trm, trm', vars) =>
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1179
                              let val (kv, kb) = K_skeleton "k" kT (Bound (length vars)) k
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1180
                              in SOME (upd $ kb $ trm, trm', kv :: vars) end
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1181
                          | NONE =>
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1182
                              let
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1183
                                val rv = ("r", rT);
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1184
                                val rb = Bound 0;
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1185
                                val (kv, kb) = K_skeleton "k" kT (Bound 1) k;
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1186
                              in SOME (upd $ kb $ rb, sel $ rb, [kv, rv]) end))
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1187
                | mk_eq_terms _ = NONE;
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1188
            in
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1189
              (case mk_eq_terms (upd $ k $ r) of
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1190
                SOME (trm, trm', vars) =>
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1191
                  SOME
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1192
                    (prove_unfold_defs thy [] []
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1193
                      (list_all (vars, Logic.mk_equals (sel $ trm, trm'))))
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1194
              | NONE => NONE)
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1195
            end
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1196
          else NONE
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15273
diff changeset
  1197
      | _ => NONE));
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
  1198
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1199
fun get_upd_acc_cong_thm upd acc thy simpset =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1200
  let
38401
c4de81b7fdec avoid variable name acc (cf. cs. 3142c1e21a0e)
haftmann
parents: 38252
diff changeset
  1201
    val insts = [("upd", cterm_of thy upd), ("ac", cterm_of thy acc)];
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1202
    val prop = Thm.concl_of (named_cterm_instantiate insts updacc_cong_triv);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1203
  in
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36173
diff changeset
  1204
    Goal.prove (ProofContext.init_global thy) [] [] prop
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1205
      (fn _ =>
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1206
        simp_tac simpset 1 THEN
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  1207
        REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac 1) THEN
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1208
        TRY (resolve_tac [updacc_cong_idI] 1))
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1209
  end;
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1210
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1211
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1212
(* upd_simproc *)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1213
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1214
(*Simplify multiple updates:
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1215
    (1) "N_update y (M_update g (N_update x (M_update f r))) =
21226
a607ae87ee81 field-update in records is generalised to take a function on the field
schirmer
parents: 21109
diff changeset
  1216
          (N_update (y o x) (M_update (g o f) r))"
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1217
    (2)  "r(|M:= M r|) = r"
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1218
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1219
  In both cases "more" updates complicate matters: for this reason
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1220
  we omit considering further updates if doing so would introduce
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1221
  both a more update and an update to a field within it.*)
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1222
val upd_simproc =
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 38544
diff changeset
  1223
  Simplifier.simproc_global @{theory HOL} "record_upd_simp" ["x"]
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1224
    (fn thy => fn _ => fn t =>
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1225
      let
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1226
        (*We can use more-updators with other updators as long
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1227
          as none of the other updators go deeper than any more
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1228
          updator. min here is the depth of the deepest other
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1229
          updator, max the depth of the shallowest more updator.*)
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1230
        fun include_depth (dep, true) (min, max) =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1231
              if min <= dep
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1232
              then SOME (min, if dep <= max orelse max = ~1 then dep else max)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1233
              else NONE
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1234
          | include_depth (dep, false) (min, max) =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1235
              if dep <= max orelse max = ~1
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1236
              then SOME (if min <= dep then dep else min, max)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1237
              else NONE;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1238
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1239
        fun getupdseq (term as (upd as Const (u, _)) $ f $ tm) min max =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1240
              (case get_update_details u thy of
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1241
                SOME (s, dep, ismore) =>
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1242
                  (case include_depth (dep, ismore) (min, max) of
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1243
                    SOME (min', max') =>
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1244
                      let val (us, bs, _) = getupdseq tm min' max'
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1245
                      in ((upd, s, f) :: us, bs, fastype_of term) end
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1246
                  | NONE => ([], term, HOLogic.unitT))
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1247
              | NONE => ([], term, HOLogic.unitT))
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1248
          | getupdseq term _ _ = ([], term, HOLogic.unitT);
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1249
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1250
        val (upds, base, baseT) = getupdseq t 0 ~1;
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1251
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1252
        fun is_upd_noop s (Abs (n, T, Const (s', T') $ tm')) tm =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1253
              if s = s' andalso null (loose_bnos tm')
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1254
                andalso subst_bound (HOLogic.unit, tm') = tm
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1255
              then (true, Abs (n, T, Const (s', T') $ Bound 1))
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1256
              else (false, HOLogic.unit)
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1257
          | is_upd_noop _ _ _ = (false, HOLogic.unit);
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1258
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1259
        fun get_noop_simps (upd as Const _) (Abs (_, _, (acc as Const _) $ _)) =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1260
          let
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1261
            val ss = get_sel_upd_defs thy;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1262
            val uathm = get_upd_acc_cong_thm upd acc thy ss;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1263
          in
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 34151
diff changeset
  1264
           [Drule.export_without_context (uathm RS updacc_noopE),
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 34151
diff changeset
  1265
            Drule.export_without_context (uathm RS updacc_noop_compE)]
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1266
          end;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1267
32974
2a1aaa2d9e64 eliminated old List.foldr and OldTerm operations;
wenzelm
parents: 32973
diff changeset
  1268
        (*If f is constant then (f o g) = f.  We know that K_skeleton
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1269
          only returns constant abstractions thus when we see an
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1270
          abstraction we can discard inner updates.*)
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1271
        fun add_upd (f as Abs _) fs = [f]
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1272
          | add_upd f fs = (f :: fs);
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1273
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1274
        (*mk_updterm returns
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1275
          (orig-term-skeleton, simplified-skeleton,
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1276
            variables, duplicate-updates, simp-flag, noop-simps)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1277
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1278
          where duplicate-updates is a table used to pass upward
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1279
          the list of update functions which can be composed
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1280
          into an update above them, simp-flag indicates whether
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1281
          any simplification was achieved, and noop-simps are
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1282
          used for eliminating case (2) defined above*)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1283
        fun mk_updterm ((upd as Const (u, T), s, f) :: upds) above term =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1284
              let
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1285
                val (lhs, rhs, vars, dups, simp, noops) =
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1286
                  mk_updterm upds (Symtab.update (u, ()) above) term;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1287
                val (fvar, skelf) =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1288
                  K_skeleton (Long_Name.base_name s) (domain_type T) (Bound (length vars)) f;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1289
                val (isnoop, skelf') = is_upd_noop s f term;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1290
                val funT = domain_type T;
35133
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1291
                fun mk_comp_local (f, f') =
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1292
                  Const (@{const_name Fun.comp}, funT --> funT --> funT) $ f $ f';
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1293
              in
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1294
                if isnoop then
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1295
                  (upd $ skelf' $ lhs, rhs, vars,
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1296
                    Symtab.update (u, []) dups, true,
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1297
                    if Symtab.defined noops u then noops
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1298
                    else Symtab.update (u, get_noop_simps upd skelf') noops)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1299
                else if Symtab.defined above u then
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1300
                  (upd $ skelf $ lhs, rhs, fvar :: vars,
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1301
                    Symtab.map_default (u, []) (add_upd skelf) dups,
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1302
                    true, noops)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1303
                else
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1304
                  (case Symtab.lookup dups u of
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1305
                    SOME fs =>
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1306
                     (upd $ skelf $ lhs,
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1307
                      upd $ foldr1 mk_comp_local (add_upd skelf fs) $ rhs,
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1308
                      fvar :: vars, dups, true, noops)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1309
                  | NONE => (upd $ skelf $ lhs, upd $ skelf $ rhs, fvar :: vars, dups, simp, noops))
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1310
              end
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1311
          | mk_updterm [] _ _ =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1312
              (Bound 0, Bound 0, [("r", baseT)], Symtab.empty, false, Symtab.empty)
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1313
          | mk_updterm us _ _ = raise TERM ("mk_updterm match", map (fn (x, _, _) => x) us);
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1314
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1315
        val (lhs, rhs, vars, _, simp, noops) = mk_updterm upds Symtab.empty base;
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32809
diff changeset
  1316
        val noops' = maps snd (Symtab.dest noops);
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1317
      in
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1318
        if simp then
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1319
          SOME
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1320
            (prove_unfold_defs thy noops' [simproc]
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1321
              (list_all (vars, Logic.mk_equals (lhs, rhs))))
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1322
        else NONE
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1323
      end);
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1324
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1325
end;
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
  1326
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1327
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1328
(* eq_simproc *)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1329
32974
2a1aaa2d9e64 eliminated old List.foldr and OldTerm operations;
wenzelm
parents: 32973
diff changeset
  1330
(*Look up the most specific record-equality.
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1331
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1332
 Note on efficiency:
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1333
 Testing equality of records boils down to the test of equality of all components.
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1334
 Therefore the complexity is: #components * complexity for single component.
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1335
 Especially if a record has a lot of components it may be better to split up
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1336
 the record first and do simplification on that (split_simp_tac).
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1337
 e.g. r(|lots of updates|) = x
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1338
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1339
             eq_simproc          split_simp_tac
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1340
 Complexity: #components * #updates     #updates
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1341
*)
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1342
val eq_simproc =
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 38544
diff changeset
  1343
  Simplifier.simproc_global @{theory HOL} "record_eq_simp" ["r = s"]
18858
ceb93f3af7f0 advanced translations: Context.generic;
wenzelm
parents: 18728
diff changeset
  1344
    (fn thy => fn _ => fn t =>
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38857
diff changeset
  1345
      (case t of Const (@{const_name HOL.eq}, Type (_, [T, _])) $ _ $ _ =>
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1346
        (case rec_id ~1 T of
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1347
          "" => NONE
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1348
        | name =>
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1349
            (case get_equalities thy name of
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1350
              NONE => NONE
35410
1ea89d2a1bd4 more antiquotations;
wenzelm
parents: 35408
diff changeset
  1351
            | SOME thm => SOME (thm RS @{thm Eq_TrueI})))
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1352
      | _ => NONE));
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1353
14079
1c22e5499eeb - record_split_tac now also works for object-level universal quantifier
berghofe
parents: 13904
diff changeset
  1354
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1355
(* split_simproc *)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1356
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1357
(*Split quantified occurrences of records, for which P holds.  P can peek on the
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1358
  subterm starting at the quantified occurrence of the record (including the quantifier):
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1359
    P t = 0: do not split
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1360
    P t = ~1: completely split
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1361
    P t > 0: split up to given bound of record extensions.*)
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1362
fun split_simproc P =
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 38544
diff changeset
  1363
  Simplifier.simproc_global @{theory HOL} "record_split_simp" ["x"]
18858
ceb93f3af7f0 advanced translations: Context.generic;
wenzelm
parents: 18728
diff changeset
  1364
    (fn thy => fn _ => fn t =>
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1365
      (case t of
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1366
        Const (quantifier, Type (_, [Type (_, [T, _]), _])) $ _ =>
35147
wenzelm
parents: 35146
diff changeset
  1367
          if quantifier = @{const_name all} orelse
wenzelm
parents: 35146
diff changeset
  1368
            quantifier = @{const_name All} orelse
35133
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1369
            quantifier = @{const_name Ex}
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1370
          then
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1371
            (case rec_id ~1 T of
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1372
              "" => NONE
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1373
            | _ =>
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1374
                let val split = P t in
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1375
                  if split <> 0 then
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1376
                    (case get_splits thy (rec_id split T) of
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1377
                      NONE => NONE
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1378
                    | SOME (all_thm, All_thm, Ex_thm, _) =>
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1379
                        SOME
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1380
                          (case quantifier of
35133
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1381
                            @{const_name all} => all_thm
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1382
                          | @{const_name All} => All_thm RS eq_reflection
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1383
                          | @{const_name Ex} => Ex_thm RS eq_reflection
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1384
                          | _ => error "split_simproc"))
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1385
                  else NONE
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1386
                end)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1387
          else NONE
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1388
      | _ => NONE));
7178
50b9849cf6ad record_simproc for sel-upd (by Sebastian Nanz);
wenzelm
parents: 6851
diff changeset
  1389
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1390
val ex_sel_eq_simproc =
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 38544
diff changeset
  1391
  Simplifier.simproc_global @{theory HOL} "ex_sel_eq_simproc" ["Ex t"]
18858
ceb93f3af7f0 advanced translations: Context.generic;
wenzelm
parents: 18728
diff changeset
  1392
    (fn thy => fn ss => fn t =>
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1393
      let
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1394
        fun prove prop =
33691
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  1395
          prove_global true thy [] prop
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1396
            (fn _ => simp_tac (Simplifier.inherit_context ss (get_simpset thy)
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1397
                addsimps @{thms simp_thms} addsimprocs [split_simproc (K ~1)]) 1);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1398
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1399
        fun mkeq (lr, Teq, (sel, Tsel), x) i =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1400
          if is_selector thy sel then
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1401
            let
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1402
              val x' =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1403
                if not (loose_bvar1 (x, 0))
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1404
                then Free ("x" ^ string_of_int i, range_type Tsel)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1405
                else raise TERM ("", [x]);
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1406
              val sel' = Const (sel, Tsel) $ Bound 0;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1407
              val (l, r) = if lr then (sel', x') else (x', sel');
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38857
diff changeset
  1408
            in Const (@{const_name HOL.eq}, Teq) $ l $ r end
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1409
          else raise TERM ("", [Const (sel, Tsel)]);
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1410
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38857
diff changeset
  1411
        fun dest_sel_eq (Const (@{const_name HOL.eq}, Teq) $ (Const (sel, Tsel) $ Bound 0) $ X) =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1412
              (true, Teq, (sel, Tsel), X)
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38857
diff changeset
  1413
          | dest_sel_eq (Const (@{const_name HOL.eq}, Teq) $ X $ (Const (sel, Tsel) $ Bound 0)) =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1414
              (false, Teq, (sel, Tsel), X)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1415
          | dest_sel_eq _ = raise TERM ("", []);
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1416
      in
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1417
        (case t of
35133
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1418
          Const (@{const_name Ex}, Tex) $ Abs (s, T, t) =>
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1419
           (let
35133
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1420
             val eq = mkeq (dest_sel_eq t) 0;
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1421
             val prop =
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1422
               list_all ([("r", T)],
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1423
                 Logic.mk_equals
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1424
                  (Const (@{const_name Ex}, Tex) $ Abs (s, T, eq), HOLogic.true_const));
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1425
            in SOME (prove prop) end
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1426
            handle TERM _ => NONE)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1427
        | _ => NONE)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1428
      end);
14427
cea7d2f76112 added record_ex_sel_eq_simproc
schirmer
parents: 14358
diff changeset
  1429
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
  1430
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1431
(* split_simp_tac *)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1432
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1433
(*Split (and simplify) all records in the goal for which P holds.
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1434
  For quantified occurrences of a record
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1435
  P can peek on the whole subterm (including the quantifier); for free variables P
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1436
  can only peek on the variable itself.
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1437
  P t = 0: do not split
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1438
  P t = ~1: completely split
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1439
  P t > 0: split up to given bound of record extensions.*)
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1440
fun split_simp_tac thms P = CSUBGOAL (fn (cgoal, i) =>
14255
e6e3e3f0deed Records:
schirmer
parents: 14079
diff changeset
  1441
  let
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1442
    val thy = Thm.theory_of_cterm cgoal;
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1443
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1444
    val goal = term_of cgoal;
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1445
    val frees = filter (is_recT o #2) (Term.add_frees goal []);
14255
e6e3e3f0deed Records:
schirmer
parents: 14079
diff changeset
  1446
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1447
    val has_rec = exists_Const
14255
e6e3e3f0deed Records:
schirmer
parents: 14079
diff changeset
  1448
      (fn (s, Type (_, [Type (_, [T, _]), _])) =>
35133
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1449
          (s = @{const_name all} orelse s = @{const_name All} orelse s = @{const_name Ex}) andalso
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  1450
          is_recT T
14255
e6e3e3f0deed Records:
schirmer
parents: 14079
diff changeset
  1451
        | _ => false);
e6e3e3f0deed Records:
schirmer
parents: 14079
diff changeset
  1452
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1453
    fun mk_split_free_tac free induct_thm i =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1454
      let
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1455
        val cfree = cterm_of thy free;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1456
        val _$ (_ $ r) = concl_of induct_thm;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1457
        val crec = cterm_of thy r;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1458
        val thm = cterm_instantiate [(crec, cfree)] induct_thm;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1459
      in
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1460
        simp_tac (HOL_basic_ss addsimps @{thms induct_atomize}) i THEN
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1461
        rtac thm i THEN
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1462
        simp_tac (HOL_basic_ss addsimps @{thms induct_rulify}) i
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1463
      end;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1464
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1465
    val split_frees_tacs =
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1466
      frees |> map_filter (fn (x, T) =>
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1467
        (case rec_id ~1 T of
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1468
          "" => NONE
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1469
        | _ =>
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1470
            let
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1471
              val free = Free (x, T);
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1472
              val split = P free;
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1473
            in
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1474
              if split <> 0 then
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1475
                (case get_splits thy (rec_id split T) of
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1476
                  NONE => NONE
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1477
                | SOME (_, _, _, induct_thm) =>
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1478
                    SOME (mk_split_free_tac free induct_thm i))
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1479
              else NONE
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1480
            end));
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1481
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1482
    val simprocs = if has_rec goal then [split_simproc P] else [];
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1483
    val thms' = K_comp_convs @ thms;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1484
  in
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1485
    EVERY split_frees_tacs THEN
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1486
    Simplifier.full_simp_tac (get_simpset thy addsimps thms' addsimprocs simprocs) i
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1487
  end);
14255
e6e3e3f0deed Records:
schirmer
parents: 14079
diff changeset
  1488
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1489
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1490
(* split_tac *)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1491
35147
wenzelm
parents: 35146
diff changeset
  1492
(*Split all records in the goal, which are quantified by !! or ALL.*)
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1493
val split_tac = CSUBGOAL (fn (cgoal, i) =>
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1494
  let
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1495
    val goal = term_of cgoal;
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1496
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1497
    val has_rec = exists_Const
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1498
      (fn (s, Type (_, [Type (_, [T, _]), _])) =>
35147
wenzelm
parents: 35146
diff changeset
  1499
          (s = @{const_name all} orelse s = @{const_name All}) andalso is_recT T
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1500
        | _ => false);
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1501
35240
wenzelm
parents: 35239
diff changeset
  1502
    fun is_all (Const (@{const_name all}, _) $ _) = ~1
wenzelm
parents: 35239
diff changeset
  1503
      | is_all (Const (@{const_name All}, _) $ _) = ~1
wenzelm
parents: 35239
diff changeset
  1504
      | is_all _ = 0;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1505
  in
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1506
    if has_rec goal then
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1507
      Simplifier.full_simp_tac (HOL_basic_ss addsimprocs [split_simproc is_all]) i
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1508
    else no_tac
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1509
  end);
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1510
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
  1511
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  1512
(* wrapper *)
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  1513
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1514
val split_name = "record_split_tac";
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1515
val split_wrapper = (split_name, fn tac => split_tac ORELSE' tac);
5698
2b5d9bdec5af field types: datatype;
wenzelm
parents: 5290
diff changeset
  1516
16330
934219e919e4 Sign.read_typ_abbrev;
wenzelm
parents: 16281
diff changeset
  1517
934219e919e4 Sign.read_typ_abbrev;
wenzelm
parents: 16281
diff changeset
  1518
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1519
(** theory extender interface **)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1520
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1521
(* prepare arguments *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1522
27278
129574589734 export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27239
diff changeset
  1523
fun read_raw_parent ctxt raw_T =
129574589734 export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27239
diff changeset
  1524
  (case ProofContext.read_typ_abbrev ctxt raw_T of
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1525
    Type (name, Ts) => (Ts, name)
27278
129574589734 export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27239
diff changeset
  1526
  | T => error ("Bad parent record specification: " ^ Syntax.string_of_typ ctxt T));
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1527
27278
129574589734 export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27239
diff changeset
  1528
fun read_typ ctxt raw_T env =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1529
  let
27278
129574589734 export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27239
diff changeset
  1530
    val ctxt' = fold (Variable.declare_typ o TFree) env ctxt;
129574589734 export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27239
diff changeset
  1531
    val T = Syntax.read_typ ctxt' raw_T;
29270
0eade173f77e moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents: 29265
diff changeset
  1532
    val env' = OldTerm.add_typ_tfrees (T, env);
27278
129574589734 export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27239
diff changeset
  1533
  in (T, env') end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1534
27278
129574589734 export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27239
diff changeset
  1535
fun cert_typ ctxt raw_T env =
129574589734 export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27239
diff changeset
  1536
  let
129574589734 export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27239
diff changeset
  1537
    val thy = ProofContext.theory_of ctxt;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1538
    val T = Type.no_tvars (Sign.certify_typ thy raw_T)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1539
      handle TYPE (msg, _, _) => error msg;
29270
0eade173f77e moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents: 29265
diff changeset
  1540
    val env' = OldTerm.add_typ_tfrees (T, env);
27278
129574589734 export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27239
diff changeset
  1541
  in (T, env') end;
129574589734 export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27239
diff changeset
  1542
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  1543
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1544
(* attributes *)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1545
33368
b1cf34f1855c modernized structure Rule_Cases;
wenzelm
parents: 33095
diff changeset
  1546
fun case_names_fields x = Rule_Cases.case_names ["fields"] x;
24830
a7b3ab44d993 moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents: 24712
diff changeset
  1547
fun induct_type_global name = [case_names_fields, Induct.induct_type name];
a7b3ab44d993 moved Pure/Isar/induct_attrib.ML and Provers/induct_method.ML to Tools/induct.ML;
wenzelm
parents: 24712
diff changeset
  1548
fun cases_type_global name = [case_names_fields, Induct.cases_type name];
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1549
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
  1550
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1551
(* tactics *)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1552
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1553
fun simp_all_tac ss simps = ALLGOALS (Simplifier.asm_full_simp_tac (ss addsimps simps));
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1554
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1555
(*Do case analysis / induction according to rule on last parameter of ith subgoal
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1556
  (or on s if there are no parameters).
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1557
  Instatiation of record variable (and predicate) in rule is calculated to
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1558
  avoid problems with higher order unification.*)
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1559
fun try_param_tac s rule = CSUBGOAL (fn (cgoal, i) =>
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1560
  let
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1561
    val cert = Thm.cterm_of (Thm.theory_of_cterm cgoal);
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1562
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1563
    val g = Thm.term_of cgoal;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1564
    val params = Logic.strip_params g;
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1565
    val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl g);
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1566
    val rule' = Thm.lift_rule cgoal rule;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1567
    val (P, ys) = strip_comb (HOLogic.dest_Trueprop
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1568
      (Logic.strip_assums_concl (prop_of rule')));
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1569
    (*ca indicates if rule is a case analysis or induction rule*)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1570
    val (x, ca) =
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
  1571
      (case rev (drop (length params) ys) of
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1572
        [] => (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1573
          (hd (rev (Logic.strip_assums_hyp (hd (prems_of rule')))))))), true)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1574
      | [x] => (head_of x, false));
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1575
    val rule'' =
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1576
      cterm_instantiate
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1577
        (map (pairself cert)
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1578
          (case rev params of
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1579
            [] =>
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1580
              (case AList.lookup (op =) (Term.add_frees g []) s of
40315
11846d9800de try_param_tac: plain user error appears more appropriate;
wenzelm
parents: 39557
diff changeset
  1581
                NONE => error "try_param_tac: no such variable"
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1582
              | SOME T => [(P, if ca then concl else lambda (Free (s, T)) concl), (x, Free (s, T))])
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1583
          | (_, T) :: _ =>
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1584
              [(P, list_abs (params, if ca then concl else incr_boundvars 1 (Abs (s, T, concl)))),
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1585
                (x, list_abs (params, Bound 0))])) rule';
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1586
  in compose_tac (false, rule'', nprems_of rule) i end);
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1587
15215
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1588
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1589
fun extension_definition name fields alphas zeta moreT more vars thy =
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1590
  let
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1591
    val base_name = Long_Name.base_name name;
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1592
32977
wenzelm
parents: 32976
diff changeset
  1593
    val fieldTs = map snd fields;
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1594
    val fields_moreTs = fieldTs @ [moreT];
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1595
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1596
    val alphas_zeta = alphas @ [zeta];
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1597
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1598
    val ext_binding = Binding.name (suffix extN base_name);
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1599
    val ext_name = suffix extN name;
38533
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1600
    val ext_tyco = suffix ext_typeN name
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1601
    val extT = Type (ext_tyco, map TFree alphas_zeta);
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1602
    val ext_type = fields_moreTs ---> extT;
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1603
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1604
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1605
    (* the tree of new types that will back the record extension *)
32767
2885e2a09f72 removed dead/duplicate code;
wenzelm
parents: 32765
diff changeset
  1606
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  1607
    val mktreeV = Balanced_Tree.make Iso_Tuple_Support.mk_cons_tuple;
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  1608
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  1609
    fun mk_iso_tuple (left, right) (thy, i) =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1610
      let
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1611
        val suff = if i = 0 then ext_typeN else inner_typeN ^ string_of_int i;
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1612
        val ((_, cons), thy') = thy
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1613
          |> Iso_Tuple_Support.add_iso_tuple_type
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1614
            (Binding.suffix_name suff (Binding.name base_name), alphas_zeta)
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1615
              (fastype_of left, fastype_of right);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1616
      in
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
  1617
        (cons $ left $ right, (thy', i + 1))
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1618
      end;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1619
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  1620
    (*trying to create a 1-element iso_tuple will fail, and is pointless anyway*)
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  1621
    fun mk_even_iso_tuple [arg] = pair arg
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  1622
      | mk_even_iso_tuple args = mk_iso_tuple (Iso_Tuple_Support.dest_cons_tuple (mktreeV args));
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1623
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1624
    fun build_meta_tree_type i thy vars more =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1625
      let val len = length vars in
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
  1626
        if len < 1 then raise TYPE ("meta_tree_type args too short", [], vars)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1627
        else if len > 16 then
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1628
          let
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1629
            fun group16 [] = []
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
  1630
              | group16 xs = take 16 xs :: group16 (drop 16 xs);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1631
            val vars' = group16 vars;
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  1632
            val (composites, (thy', i')) = fold_map mk_even_iso_tuple vars' (thy, i);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1633
          in
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1634
            build_meta_tree_type i' thy' composites more
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1635
          end
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1636
        else
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  1637
          let val (term, (thy', _)) = mk_iso_tuple (mktreeV vars, more) (thy, 0)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1638
          in (term, thy') end
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1639
      end;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1640
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1641
    val _ = timing_msg "record extension preparing definitions";
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1642
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1643
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1644
    (* 1st stage part 1: introduce the tree of new types *)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1645
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1646
    fun get_meta_tree () = build_meta_tree_type 1 thy vars more;
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1647
    val (ext_body, typ_thy) =
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1648
      timeit_msg "record extension nested type def:" get_meta_tree;
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1649
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1650
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1651
    (* prepare declarations and definitions *)
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1652
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1653
    (* 1st stage part 2: define the ext constant *)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1654
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1655
    fun mk_ext args = list_comb (Const (ext_name, ext_type), args);
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1656
    val ext_spec = Logic.mk_equals (mk_ext (vars @ [more]), ext_body);
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1657
16379
d29d27e0f59f more timing information
schirmer
parents: 16364
diff changeset
  1658
    fun mk_defs () =
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1659
      typ_thy
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1660
      |> Sign.declare_const ((ext_binding, ext_type), NoSyn) |> snd
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39134
diff changeset
  1661
      |> Global_Theory.add_defs false [((Thm.def_binding ext_binding, ext_spec), [])]
33691
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  1662
      ||> Theory.checkpoint
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1663
    val ([ext_def], defs_thy) =
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1664
      timeit_msg "record extension constructor def:" mk_defs;
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1665
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1666
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1667
    (* prepare propositions *)
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1668
16379
d29d27e0f59f more timing information
schirmer
parents: 16364
diff changeset
  1669
    val _ = timing_msg "record extension preparing propositions";
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1670
    val vars_more = vars @ [more];
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1671
    val variants = map (fn Free (x, _) => x) vars_more;
15215
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1672
    val ext = mk_ext vars_more;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1673
    val s = Free (rN, extT);
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
  1674
    val P = Free (Name.variant variants "P", extT --> HOLogic.boolT);
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1675
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1676
    val inject_prop =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1677
      let val vars_more' = map (fn (Free (x, T)) => Free (x ^ "'", T)) vars_more in
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1678
        HOLogic.mk_conj (HOLogic.eq_const extT $
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1679
          mk_ext vars_more $ mk_ext vars_more', HOLogic.true_const)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1680
        ===
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1681
        foldr1 HOLogic.mk_conj
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1682
          (map HOLogic.mk_eq (vars_more ~~ vars_more') @ [HOLogic.true_const])
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1683
      end;
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1684
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1685
    val induct_prop =
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  1686
      (All (map dest_Free vars_more) (Trueprop (P $ ext)), Trueprop (P $ s));
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  1687
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
  1688
    val split_meta_prop =
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1689
      let val P = Free (Name.variant variants "P", extT --> Term.propT) in
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1690
        Logic.mk_equals
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
  1691
         (All [dest_Free s] (P $ s), All (map dest_Free vars_more) (P $ ext))
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1692
      end;
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
  1693
33691
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  1694
    val prove_standard = prove_future_global true defs_thy;
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1695
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1696
    fun inject_prf () =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1697
      simplify HOL_ss
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1698
        (prove_standard [] inject_prop
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1699
          (fn _ =>
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1700
            simp_tac (HOL_basic_ss addsimps [ext_def]) 1 THEN
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1701
            REPEAT_DETERM
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1702
              (rtac refl_conj_eq 1 ORELSE
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  1703
                Iso_Tuple_Support.iso_tuple_intros_tac 1 ORELSE
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1704
                rtac refl 1)));
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1705
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  1706
    val inject = timeit_msg "record extension inject proof:" inject_prf;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1707
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1708
    (*We need a surjection property r = (| f = f r, g = g r ... |)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1709
      to prove other theorems. We haven't given names to the accessors
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1710
      f, g etc yet however, so we generate an ext structure with
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1711
      free variables as all arguments and allow the introduction tactic to
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 34151
diff changeset
  1712
      operate on it as far as it can. We then use Drule.export_without_context
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 34151
diff changeset
  1713
      to convert the free variables into unifiable variables and unify them with
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1714
      (roughly) the definition of the accessor.*)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1715
    fun surject_prf () =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1716
      let
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1717
        val cterm_ext = cterm_of defs_thy ext;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1718
        val start = named_cterm_instantiate [("y", cterm_ext)] surject_assist_idE;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1719
        val tactic1 =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1720
          simp_tac (HOL_basic_ss addsimps [ext_def]) 1 THEN
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  1721
          REPEAT_ALL_NEW Iso_Tuple_Support.iso_tuple_intros_tac 1;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1722
        val tactic2 = REPEAT (rtac surject_assistI 1 THEN rtac refl 1);
32972
45ba8b02e1e4 misc tuning and simplification;
wenzelm
parents: 32970
diff changeset
  1723
        val [halfway] = Seq.list_of (tactic1 start);
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 34151
diff changeset
  1724
        val [surject] = Seq.list_of (tactic2 (Drule.export_without_context halfway));
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1725
      in
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1726
        surject
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1727
      end;
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1728
    val surject = timeit_msg "record extension surjective proof:" surject_prf;
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1729
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1730
    fun split_meta_prf () =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1731
      prove_standard [] split_meta_prop
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1732
        (fn _ =>
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1733
          EVERY1
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1734
           [rtac equal_intr_rule, Goal.norm_hhf_tac,
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1735
            etac meta_allE, atac,
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1736
            rtac (prop_subst OF [surject]),
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1737
            REPEAT o etac meta_allE, atac]);
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1738
    val split_meta = timeit_msg "record extension split_meta proof:" split_meta_prf;
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  1739
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  1740
    fun induct_prf () =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1741
      let val (assm, concl) = induct_prop in
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1742
        prove_standard [assm] concl
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1743
          (fn {prems, ...} =>
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1744
            cut_rules_tac [split_meta RS Drule.equal_elim_rule2] 1 THEN
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1745
            resolve_tac prems 2 THEN
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  1746
            asm_simp_tac HOL_ss 1)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1747
      end;
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  1748
    val induct = timeit_msg "record extension induct proof:" induct_prf;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1749
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  1750
    val ([induct', inject', surjective', split_meta'], thm_thy) =
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1751
      defs_thy
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39134
diff changeset
  1752
      |> Global_Theory.add_thms (map (Thm.no_attributes o apfst Binding.name)
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  1753
           [("ext_induct", induct),
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  1754
            ("ext_inject", inject),
32744
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
  1755
            ("ext_surjective", surject),
38533
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1756
            ("ext_split", split_meta)]);
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1757
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1758
  in
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1759
    (((ext_name, ext_type), (ext_tyco, alphas_zeta),
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1760
      extT, induct', inject', surjective', split_meta', ext_def), thm_thy)
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1761
  end;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1762
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1763
fun chunks [] [] = []
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1764
  | chunks [] xs = [xs]
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
  1765
  | chunks (l :: ls) xs = take l xs :: chunks ls (drop l xs);
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1766
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
  1767
fun chop_last [] = error "chop_last: list should not be empty"
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1768
  | chop_last [x] = ([], x)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1769
  | chop_last (x :: xs) = let val (tl, l) = chop_last xs in (x :: tl, l) end;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1770
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1771
fun subst_last _ [] = error "subst_last: list should not be empty"
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  1772
  | subst_last s [_] = [s]
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1773
  | subst_last s (x :: xs) = x :: subst_last s xs;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1774
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1775
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1776
(* mk_recordT *)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1777
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1778
(*build up the record type from the current extension tpye extT and a list
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1779
  of parent extensions, starting with the root of the record hierarchy*)
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
  1780
fun mk_recordT extT =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1781
  fold_rev (fn (parent, Ts) => fn T => Type (parent, subst_last T Ts)) extT;
15215
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1782
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1783
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1784
fun obj_to_meta_all thm =
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1785
  let
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1786
    fun E thm =  (* FIXME proper name *)
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1787
      (case SOME (spec OF [thm]) handle THM _ => NONE of
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1788
        SOME thm' => E thm'
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1789
      | NONE => thm);
15215
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1790
    val th1 = E thm;
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1791
    val th2 = Drule.forall_intr_vars th1;
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1792
  in th2 end;
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1793
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1794
fun meta_to_obj_all thm =
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1795
  let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26496
diff changeset
  1796
    val thy = Thm.theory_of_thm thm;
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26496
diff changeset
  1797
    val prop = Thm.prop_of thm;
15215
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1798
    val params = Logic.strip_params prop;
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1799
    val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl prop);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1800
    val ct = cterm_of thy (HOLogic.mk_Trueprop (HOLogic.list_all (params, concl)));
15215
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1801
    val thm' = Seq.hd (REPEAT (rtac allI 1) (Thm.trivial ct));
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1802
  in Thm.implies_elim thm' thm end;
15215
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1803
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  1804
38533
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1805
(* code generation *)
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1806
38544
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1807
fun instantiate_random_record tyco vs extN Ts thy =
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1808
  let
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1809
    val size = @{term "i::code_numeral"};
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1810
    fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit => term"});
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1811
    val T = Type (tyco, map TFree vs);
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1812
    val Tm = termifyT T;
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1813
    val params = Name.names Name.context "x" Ts;
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1814
    val lhs = HOLogic.mk_random T size;
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1815
    val tc = HOLogic.mk_return Tm @{typ Random.seed}
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1816
      (HOLogic.mk_valtermify_app extN params T);
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1817
    val rhs =
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1818
      HOLogic.mk_ST
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1819
        (map (fn (v, T') =>
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1820
          ((HOLogic.mk_random T' size, @{typ Random.seed}), SOME (v, termifyT T'))) params)
38544
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1821
        tc @{typ Random.seed} (SOME Tm, @{typ Random.seed});
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1822
    val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1823
  in 
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1824
    thy
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1825
    |> Class.instantiation ([tyco], vs, @{sort random})
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1826
    |> `(fn lthy => Syntax.check_term lthy eq)
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1827
    |-> (fn eq => Specification.definition (NONE, (apfst Binding.conceal Attrib.empty_binding, eq)))
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1828
    |> snd
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1829
    |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1830
  end;
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1831
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1832
fun ensure_random_record ext_tyco vs extN Ts thy =
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1833
  let
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1834
    val algebra = Sign.classes_of thy;
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1835
    val has_inst = can (Sorts.mg_domain algebra ext_tyco) @{sort random};
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1836
  in if has_inst then thy
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1837
    else case Quickcheck_Generators.perhaps_constrain thy (map (rpair @{sort random}) Ts) vs
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1838
     of SOME constrain => instantiate_random_record ext_tyco (map constrain vs) extN
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1839
          ((map o map_atyps) (fn TFree v => TFree (constrain v)) Ts) thy
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1840
      | NONE => thy
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1841
  end;
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1842
38533
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1843
fun add_code ext_tyco vs extT ext simps inject thy =
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1844
  let
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1845
    val eq =
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1846
      (HOLogic.mk_Trueprop o HOLogic.mk_eq)
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38758
diff changeset
  1847
        (Const (@{const_name HOL.equal}, extT --> extT --> HOLogic.boolT),
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38857
diff changeset
  1848
         Const (@{const_name HOL.eq}, extT --> extT --> HOLogic.boolT));
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1849
    fun tac eq_def =
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1850
      Class.intro_classes_tac []
38533
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1851
      THEN (Simplifier.rewrite_goals_tac [Simpdata.mk_eq eq_def])
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1852
      THEN ALLGOALS (rtac @{thm refl});
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1853
    fun mk_eq thy eq_def = Simplifier.rewrite_rule
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1854
      [(AxClass.unoverload thy o Thm.symmetric o Simpdata.mk_eq) eq_def] inject;
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1855
    fun mk_eq_refl thy =
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38758
diff changeset
  1856
      @{thm equal_refl}
38533
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1857
      |> Thm.instantiate
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38758
diff changeset
  1858
        ([pairself (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort equal}), Logic.varifyT_global extT)], [])
38533
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1859
      |> AxClass.unoverload thy;
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1860
  in
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1861
    thy
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1862
    |> Code.add_datatype [ext]
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1863
    |> fold Code.add_default_eqn simps
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38758
diff changeset
  1864
    |> Class.instantiation ([ext_tyco], vs, [HOLogic.class_equal])
38533
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1865
    |> `(fn lthy => Syntax.check_term lthy eq)
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1866
    |-> (fn eq => Specification.definition
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1867
         (NONE, (Attrib.empty_binding, eq)))
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1868
    |-> (fn (_, (_, eq_def)) =>
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1869
       Class.prove_instantiation_exit_result Morphism.thm
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1870
         (fn _ => fn eq_def => tac eq_def) eq_def)
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1871
    |-> (fn eq_def => fn thy => thy |> Code.del_eqn eq_def |> Code.add_default_eqn (mk_eq thy eq_def))
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1872
    |> (fn thy => Code.add_nbe_default_eqn (mk_eq_refl thy) thy)
38544
ac554311b1b9 re-added instantiation of type class random for records
haftmann
parents: 38534
diff changeset
  1873
    |> ensure_random_record ext_tyco vs (fst ext) ((fst o strip_type o snd) ext)
38533
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1874
  end;
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1875
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  1876
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1877
(* definition *)
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1878
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  1879
fun definition (alphas, binding) parent (parents: parent_info list) raw_fields thy =
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1880
  let
37470
fcc768dc9dd0 more binding; avoid arcane Rep and Abs prefixes
haftmann
parents: 37469
diff changeset
  1881
    val prefix = Binding.name_of binding;
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1882
    val name = Sign.full_name thy binding;
37470
fcc768dc9dd0 more binding; avoid arcane Rep and Abs prefixes
haftmann
parents: 37469
diff changeset
  1883
    val full = Sign.full_name_path thy prefix;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1884
35136
34206672b168 modernized signature -- proper binding;
wenzelm
parents: 35135
diff changeset
  1885
    val bfields = map (fn (x, T, _) => (x, T)) raw_fields;
34206672b168 modernized signature -- proper binding;
wenzelm
parents: 35135
diff changeset
  1886
    val field_syntax = map #3 raw_fields;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1887
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32809
diff changeset
  1888
    val parent_fields = maps #fields parents;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1889
    val parent_chunks = map (length o #fields) parents;
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1890
    val parent_names = map fst parent_fields;
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1891
    val parent_types = map snd parent_fields;
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1892
    val parent_fields_len = length parent_fields;
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1893
    val parent_variants =
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1894
      Name.variant_list [moreN, rN, rN ^ "'", wN] (map Long_Name.base_name parent_names);
37470
fcc768dc9dd0 more binding; avoid arcane Rep and Abs prefixes
haftmann
parents: 37469
diff changeset
  1895
    val parent_vars = map2 (curry Free) parent_variants parent_types;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1896
    val parent_len = length parents;
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1897
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1898
    val fields = map (apfst full) bfields;
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1899
    val names = map fst fields;
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1900
    val types = map snd fields;
36151
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  1901
    val alphas_fields = fold Term.add_tfreesT types [];
33049
c38f02fdf35d curried inter as canonical list operation (beware of argument order)
haftmann
parents: 33039
diff changeset
  1902
    val alphas_ext = inter (op =) alphas_fields alphas;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1903
    val len = length fields;
30715
e23e15f52d42 avoid mixing of left/right associative infixes, to make it work with experimental Poly/ML 5.3 branch;
wenzelm
parents: 30364
diff changeset
  1904
    val variants =
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
  1905
      Name.variant_list (moreN :: rN :: (rN ^ "'") :: wN :: parent_variants)
35136
34206672b168 modernized signature -- proper binding;
wenzelm
parents: 35135
diff changeset
  1906
        (map (Binding.name_of o fst) bfields);
37470
fcc768dc9dd0 more binding; avoid arcane Rep and Abs prefixes
haftmann
parents: 37469
diff changeset
  1907
    val vars = map2 (curry Free) variants types;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1908
    val named_vars = names ~~ vars;
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1909
    val idxms = 0 upto len;
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1910
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1911
    val all_fields = parent_fields @ fields;
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1912
    val all_types = parent_types @ types;
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1913
    val all_variants = parent_variants @ variants;
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1914
    val all_vars = parent_vars @ vars;
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1915
    val all_named_vars = (parent_names ~~ parent_vars) @ named_vars;
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1916
36151
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  1917
    val zeta = (Name.variant (map #1 alphas) "'z", HOLogic.typeS);
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  1918
    val moreT = TFree zeta;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1919
    val more = Free (moreN, moreT);
35136
34206672b168 modernized signature -- proper binding;
wenzelm
parents: 35135
diff changeset
  1920
    val full_moreN = full (Binding.name moreN);
34206672b168 modernized signature -- proper binding;
wenzelm
parents: 35135
diff changeset
  1921
    val bfields_more = bfields @ [(Binding.name moreN, moreT)];
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1922
    val fields_more = fields @ [(full_moreN, moreT)];
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1923
    val named_vars_more = named_vars @ [(full_moreN, more)];
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1924
    val all_vars_more = all_vars @ [more];
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1925
    val all_named_vars_more = all_named_vars @ [(full_moreN, more)];
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1926
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1927
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1928
    (* 1st stage: ext_thy *)
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1929
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1930
    val extension_name = full binding;
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1931
38758
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1932
    val ((ext, (ext_tyco, vs),
f2cfb2cc03e8 misc tuning and simplification, notably theory data;
wenzelm
parents: 38715
diff changeset
  1933
        extT, ext_induct, ext_inject, ext_surjective, ext_split, ext_def), ext_thy) =
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1934
      thy
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1935
      |> Sign.qualified_path false binding
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1936
      |> extension_definition extension_name fields alphas_ext zeta moreT more vars;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1937
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1938
    val _ = timing_msg "record preparing definitions";
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1939
    val Type extension_scheme = extT;
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1940
    val extension_name = unsuffix ext_typeN (fst extension_scheme);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1941
    val extension = let val (n, Ts) = extension_scheme in (n, subst_last HOLogic.unitT Ts) end;
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1942
    val extension_names = map (unsuffix ext_typeN o fst o #extension) parents @ [extension_name];
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
  1943
    val extension_id = implode extension_names;
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1944
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
  1945
    fun rec_schemeT n = mk_recordT (map #extension (drop n parents)) extT;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1946
    val rec_schemeT0 = rec_schemeT 0;
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1947
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1948
    fun recT n =
32972
45ba8b02e1e4 misc tuning and simplification;
wenzelm
parents: 32970
diff changeset
  1949
      let val (c, Ts) = extension in
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
  1950
        mk_recordT (map #extension (drop n parents))
32972
45ba8b02e1e4 misc tuning and simplification;
wenzelm
parents: 32970
diff changeset
  1951
          (Type (c, subst_last HOLogic.unitT Ts))
45ba8b02e1e4 misc tuning and simplification;
wenzelm
parents: 32970
diff changeset
  1952
      end;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1953
    val recT0 = recT 0;
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1954
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1955
    fun mk_rec args n =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1956
      let
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1957
        val (args', more) = chop_last args;
32974
2a1aaa2d9e64 eliminated old List.foldr and OldTerm operations;
wenzelm
parents: 32973
diff changeset
  1958
        fun mk_ext' ((name, T), args) more = mk_ext (name, T) (args @ [more]);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1959
        fun build Ts =
35430
df2862dc23a8 adapted to authentic syntax -- actual types are verbatim;
wenzelm
parents: 35410
diff changeset
  1960
          fold_rev mk_ext' (drop n ((extension_names ~~ Ts) ~~ chunks parent_chunks args')) more;
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1961
      in
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1962
        if more = HOLogic.unit
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33055
diff changeset
  1963
        then build (map_range recT (parent_len + 1))
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33055
diff changeset
  1964
        else build (map_range rec_schemeT (parent_len + 1))
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1965
      end;
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1966
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1967
    val r_rec0 = mk_rec all_vars_more 0;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1968
    val r_rec_unit0 = mk_rec (all_vars @ [HOLogic.unit]) 0;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1969
35430
df2862dc23a8 adapted to authentic syntax -- actual types are verbatim;
wenzelm
parents: 35410
diff changeset
  1970
    fun r n = Free (rN, rec_schemeT n);
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1971
    val r0 = r 0;
35430
df2862dc23a8 adapted to authentic syntax -- actual types are verbatim;
wenzelm
parents: 35410
diff changeset
  1972
    fun r_unit n = Free (rN, recT n);
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1973
    val r_unit0 = r_unit 0;
35430
df2862dc23a8 adapted to authentic syntax -- actual types are verbatim;
wenzelm
parents: 35410
diff changeset
  1974
    val w = Free (wN, rec_schemeT 0);
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1975
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1976
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1977
    (* print translations *)
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1978
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
  1979
    val record_ext_type_abbr_tr's =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1980
      let
35430
df2862dc23a8 adapted to authentic syntax -- actual types are verbatim;
wenzelm
parents: 35410
diff changeset
  1981
        val trname = hd extension_names;
35148
3a34fee2cfcd eliminated camel case;
wenzelm
parents: 35147
diff changeset
  1982
        val last_ext = unsuffix ext_typeN (fst extension);
35430
df2862dc23a8 adapted to authentic syntax -- actual types are verbatim;
wenzelm
parents: 35410
diff changeset
  1983
      in [record_ext_type_abbr_tr' name alphas zeta last_ext rec_schemeT0 trname] end;
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
  1984
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
  1985
    val record_ext_type_tr's =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  1986
      let
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
  1987
        (*avoid conflict with record_type_abbr_tr's*)
35430
df2862dc23a8 adapted to authentic syntax -- actual types are verbatim;
wenzelm
parents: 35410
diff changeset
  1988
        val trnames = if parent_len > 0 then [extension_name] else [];
35149
eee63670b5aa simplified/clarified record print translations;
wenzelm
parents: 35148
diff changeset
  1989
      in map record_ext_type_tr' trnames end;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1990
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1991
    val advanced_print_translation =
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1992
      map field_update_tr' (full_moreN :: names) @ [record_ext_tr' extension_name] @
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1993
      record_ext_type_tr's @ record_ext_type_abbr_tr's;
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  1994
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  1995
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1996
    (* prepare declarations *)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  1997
35136
34206672b168 modernized signature -- proper binding;
wenzelm
parents: 35135
diff changeset
  1998
    val sel_decls = map (mk_selC rec_schemeT0 o apfst Binding.name_of) bfields_more;
34206672b168 modernized signature -- proper binding;
wenzelm
parents: 35135
diff changeset
  1999
    val upd_decls = map (mk_updC updateN rec_schemeT0 o apfst Binding.name_of) bfields_more;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2000
    val make_decl = (makeN, all_types ---> recT0);
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2001
    val fields_decl = (fields_selN, types ---> Type extension);
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2002
    val extend_decl = (extendN, recT0 --> moreT --> rec_schemeT0);
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2003
    val truncate_decl = (truncateN, rec_schemeT0 --> recT0);
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2004
35133
a68e4972fd31 formal markup of constants;
wenzelm
parents: 35131
diff changeset
  2005
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2006
    (* prepare definitions *)
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2007
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  2008
    val ext_defs = ext_def :: map #ext_def parents;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2009
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  2010
    (*Theorems from the iso_tuple intros.
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2011
      By unfolding ext_defs from r_rec0 we create a tree of constructor
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2012
      calls (many of them Pair, but others as well). The introduction
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2013
      rules for update_accessor_eq_assist can unify two different ways
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2014
      on these constructors. If we take the complete result sequence of
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2015
      running a the introduction tactic, we get one theorem for each upd/acc
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2016
      pair, from which we can derive the bodies of our selector and
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2017
      updator and their convs.*)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2018
    fun get_access_update_thms () =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2019
      let
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2020
        val r_rec0_Vars =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2021
          let
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2022
            (*pick variable indices of 1 to avoid possible variable
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2023
              collisions with existing variables in updacc_eq_triv*)
32757
4e97fc468a53 Avoid a possible variable name conflict in instantiating a theorem.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32752
diff changeset
  2024
            fun to_Var (Free (c, T)) = Var ((c, 1), T);
4e97fc468a53 Avoid a possible variable name conflict in instantiating a theorem.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32752
diff changeset
  2025
          in mk_rec (map to_Var all_vars_more) 0 end;
4e97fc468a53 Avoid a possible variable name conflict in instantiating a theorem.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32752
diff changeset
  2026
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2027
        val cterm_rec = cterm_of ext_thy r_rec0;
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2028
        val cterm_vrs = cterm_of ext_thy r_rec0_Vars;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2029
        val insts = [("v", cterm_rec), ("v'", cterm_vrs)];
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2030
        val init_thm = named_cterm_instantiate insts updacc_eq_triv;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2031
        val terminal = rtac updacc_eq_idI 1 THEN rtac refl 1;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2032
        val tactic =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2033
          simp_tac (HOL_basic_ss addsimps ext_defs) 1 THEN
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  2034
          REPEAT (Iso_Tuple_Support.iso_tuple_intros_tac 1 ORELSE terminal);
32972
45ba8b02e1e4 misc tuning and simplification;
wenzelm
parents: 32970
diff changeset
  2035
        val updaccs = Seq.list_of (tactic init_thm);
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2036
      in
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2037
        (updaccs RL [updacc_accessor_eqE],
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2038
         updaccs RL [updacc_updator_eqE],
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2039
         updaccs RL [updacc_cong_from_eq])
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2040
      end;
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2041
    val (accessor_thms, updator_thms, upd_acc_cong_assists) =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2042
      timeit_msg "record getting tree access/updates:" get_access_update_thms;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2043
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
  2044
    fun lastN xs = drop parent_fields_len xs;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2045
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2046
    (*selectors*)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2047
    fun mk_sel_spec ((c, T), thm) =
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2048
      let
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2049
        val (acc $ arg, _) =
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2050
          HOLogic.dest_eq (HOLogic.dest_Trueprop (Envir.beta_eta_contract (Thm.concl_of thm)));
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2051
        val _ =
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2052
          if arg aconv r_rec0 then ()
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2053
          else raise TERM ("mk_sel_spec: different arg", [arg]);
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2054
      in
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2055
        Const (mk_selC rec_schemeT0 (c, T)) :== acc
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2056
      end;
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2057
    val sel_specs = map mk_sel_spec (fields_more ~~ lastN accessor_thms);
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2058
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2059
    (*updates*)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2060
    fun mk_upd_spec ((c, T), thm) =
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2061
      let
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2062
        val (upd $ _ $ arg, _) =
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2063
          HOLogic.dest_eq (HOLogic.dest_Trueprop (Envir.beta_eta_contract (Thm.concl_of thm)));
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2064
        val _ =
35135
1667fd3b051d tuned errors;
wenzelm
parents: 35133
diff changeset
  2065
          if arg aconv r_rec0 then ()
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2066
          else raise TERM ("mk_sel_spec: different arg", [arg]);
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2067
      in Const (mk_updC updateN rec_schemeT0 (c, T)) :== upd end;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2068
    val upd_specs = map mk_upd_spec (fields_more ~~ lastN updator_thms);
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2069
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2070
    (*derived operations*)
35144
wenzelm
parents: 35142
diff changeset
  2071
    val make_spec =
wenzelm
parents: 35142
diff changeset
  2072
      list_comb (Const (full (Binding.name makeN), all_types ---> recT0), all_vars) :==
wenzelm
parents: 35142
diff changeset
  2073
        mk_rec (all_vars @ [HOLogic.unit]) 0;
wenzelm
parents: 35142
diff changeset
  2074
    val fields_spec =
wenzelm
parents: 35142
diff changeset
  2075
      list_comb (Const (full (Binding.name fields_selN), types ---> Type extension), vars) :==
wenzelm
parents: 35142
diff changeset
  2076
        mk_rec (all_vars @ [HOLogic.unit]) parent_len;
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2077
    val extend_spec =
35136
34206672b168 modernized signature -- proper binding;
wenzelm
parents: 35135
diff changeset
  2078
      Const (full (Binding.name extendN), recT0 --> moreT --> rec_schemeT0) $ r_unit0 $ more :==
35144
wenzelm
parents: 35142
diff changeset
  2079
        mk_rec ((map (mk_sel r_unit0) all_fields) @ [more]) 0;
wenzelm
parents: 35142
diff changeset
  2080
    val truncate_spec =
wenzelm
parents: 35142
diff changeset
  2081
      Const (full (Binding.name truncateN), rec_schemeT0 --> recT0) $ r0 :==
wenzelm
parents: 35142
diff changeset
  2082
        mk_rec ((map (mk_sel r0) all_fields) @ [HOLogic.unit]) 0;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2083
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2084
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2085
    (* 2st stage: defs_thy *)
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2086
16379
d29d27e0f59f more timing information
schirmer
parents: 16364
diff changeset
  2087
    fun mk_defs () =
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2088
      ext_thy
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2089
      |> Sign.add_advanced_trfuns ([], [], advanced_print_translation, [])
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2090
      |> Sign.restore_naming thy
36173
99212848c933 modernized type abbreviations;
wenzelm
parents: 36159
diff changeset
  2091
      |> Typedecl.abbrev_global (binding, map #1 alphas, NoSyn) recT0 |> snd
99212848c933 modernized type abbreviations;
wenzelm
parents: 36159
diff changeset
  2092
      |> Typedecl.abbrev_global
99212848c933 modernized type abbreviations;
wenzelm
parents: 36159
diff changeset
  2093
        (Binding.suffix_name schemeN binding, map #1 (alphas @ [zeta]), NoSyn) rec_schemeT0 |> snd
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2094
      |> Sign.qualified_path false binding
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2095
      |> fold (fn ((x, T), mx) => snd o Sign.declare_const ((Binding.name x, T), mx))
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2096
        (sel_decls ~~ (field_syntax @ [NoSyn]))
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2097
      |> fold (fn (x, T) => snd o Sign.declare_const ((Binding.name x, T), NoSyn))
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2098
        (upd_decls @ [make_decl, fields_decl, extend_decl, truncate_decl])
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39134
diff changeset
  2099
      |> (Global_Theory.add_defs false o
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39134
diff changeset
  2100
            map (Thm.no_attributes o apfst (Binding.conceal o Binding.name))) sel_specs
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39134
diff changeset
  2101
      ||>> (Global_Theory.add_defs false o
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39134
diff changeset
  2102
            map (Thm.no_attributes o apfst (Binding.conceal o Binding.name))) upd_specs
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39134
diff changeset
  2103
      ||>> (Global_Theory.add_defs false o
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39134
diff changeset
  2104
            map (Thm.no_attributes o apfst (Binding.conceal o Binding.name)))
35142
495c623f1e3c conceal internal record definitions;
wenzelm
parents: 35138
diff changeset
  2105
        [make_spec, fields_spec, extend_spec, truncate_spec]
33691
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  2106
      ||> Theory.checkpoint
22747
0c9c413b4678 add definitions explicitly to code generator table
haftmann
parents: 22693
diff changeset
  2107
    val (((sel_defs, upd_defs), derived_defs), defs_thy) =
0c9c413b4678 add definitions explicitly to code generator table
haftmann
parents: 22693
diff changeset
  2108
      timeit_msg "record trfuns/tyabbrs/selectors/updates/make/fields/extend/truncate defs:"
0c9c413b4678 add definitions explicitly to code generator table
haftmann
parents: 22693
diff changeset
  2109
        mk_defs;
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2110
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2111
    (* prepare propositions *)
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2112
    val _ = timing_msg "record preparing propositions";
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2113
    val P = Free (Name.variant all_variants "P", rec_schemeT0 --> HOLogic.boolT);
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 20049
diff changeset
  2114
    val C = Free (Name.variant all_variants "C", HOLogic.boolT);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2115
    val P_unit = Free (Name.variant all_variants "P", recT0 --> HOLogic.boolT);
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2116
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2117
    (*selectors*)
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2118
    val sel_conv_props =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2119
       map (fn (c, x as Free (_, T)) => mk_sel r_rec0 (c, T) === x) named_vars_more;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2120
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2121
    (*updates*)
37470
fcc768dc9dd0 more binding; avoid arcane Rep and Abs prefixes
haftmann
parents: 37469
diff changeset
  2122
    fun mk_upd_prop i (c, T) =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2123
      let
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2124
        val x' = Free (Name.variant all_variants (Long_Name.base_name c ^ "'"), T --> T);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2125
        val n = parent_fields_len + i;
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2126
        val args' = nth_map n (K (x' $ nth all_vars_more n)) all_vars_more;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2127
      in mk_upd updateN c x' r_rec0 === mk_rec args' 0 end;
37470
fcc768dc9dd0 more binding; avoid arcane Rep and Abs prefixes
haftmann
parents: 37469
diff changeset
  2128
    val upd_conv_props = map2 mk_upd_prop idxms fields_more;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2129
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2130
    (*induct*)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2131
    val induct_scheme_prop =
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2132
      All (map dest_Free all_vars_more) (Trueprop (P $ r_rec0)) ==> Trueprop (P $ r0);
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2133
    val induct_prop =
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2134
      (All (map dest_Free all_vars) (Trueprop (P_unit $ r_rec_unit0)),
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2135
        Trueprop (P_unit $ r_unit0));
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2136
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2137
    (*surjective*)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2138
    val surjective_prop =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2139
      let val args = map (fn (c, Free (_, T)) => mk_sel r0 (c, T)) all_named_vars_more
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2140
      in r0 === mk_rec args 0 end;
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2141
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2142
    (*cases*)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2143
    val cases_scheme_prop =
32809
e72347dd3e64 avoid mixed l/r infixes, which do not work in some versions of SML;
wenzelm
parents: 32808
diff changeset
  2144
      (All (map dest_Free all_vars_more) ((r0 === r_rec0) ==> Trueprop C))
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  2145
        ==> Trueprop C;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2146
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2147
    val cases_prop =
32809
e72347dd3e64 avoid mixed l/r infixes, which do not work in some versions of SML;
wenzelm
parents: 32808
diff changeset
  2148
      (All (map dest_Free all_vars) ((r_unit0 === r_rec_unit0) ==> Trueprop C))
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  2149
         ==> Trueprop C;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2150
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2151
    (*split*)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2152
    val split_meta_prop =
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 20049
diff changeset
  2153
      let val P = Free (Name.variant all_variants "P", rec_schemeT0-->Term.propT) in
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2154
        Logic.mk_equals
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2155
         (All [dest_Free r0] (P $ r0), All (map dest_Free all_vars_more) (P $ r_rec0))
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2156
      end;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2157
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2158
    val split_object_prop =
32974
2a1aaa2d9e64 eliminated old List.foldr and OldTerm operations;
wenzelm
parents: 32973
diff changeset
  2159
      let val ALL = fold_rev (fn (v, T) => fn t => HOLogic.mk_all (v, T, t))
2a1aaa2d9e64 eliminated old List.foldr and OldTerm operations;
wenzelm
parents: 32973
diff changeset
  2160
      in ALL [dest_Free r0] (P $ r0) === ALL (map dest_Free all_vars_more) (P $ r_rec0) end;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2161
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2162
    val split_ex_prop =
32974
2a1aaa2d9e64 eliminated old List.foldr and OldTerm operations;
wenzelm
parents: 32973
diff changeset
  2163
      let val EX = fold_rev (fn (v, T) => fn t => HOLogic.mk_exists (v, T, t))
2a1aaa2d9e64 eliminated old List.foldr and OldTerm operations;
wenzelm
parents: 32973
diff changeset
  2164
      in EX [dest_Free r0] (P $ r0) === EX (map dest_Free all_vars_more) (P $ r_rec0) end;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2165
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2166
    (*equality*)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2167
    val equality_prop =
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2168
      let
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2169
        val s' = Free (rN ^ "'", rec_schemeT0);
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2170
        fun mk_sel_eq (c, Free (_, T)) = mk_sel r0 (c, T) === mk_sel s' (c, T);
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2171
        val seleqs = map mk_sel_eq all_named_vars_more;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2172
      in All (map dest_Free [r0, s']) (Logic.list_implies (seleqs, r0 === s')) end;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2173
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2174
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2175
    (* 3rd stage: thms_thy *)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2176
33691
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  2177
    fun prove stndrd = prove_future_global stndrd defs_thy;
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  2178
    val prove_standard = prove_future_global true defs_thy;
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  2179
    val future_forward_prf = future_forward_prf_standard defs_thy;
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2180
15215
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  2181
    fun prove_simp stndrd ss simps =
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2182
      let val tac = simp_all_tac ss simps
15215
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  2183
      in fn prop => prove stndrd [] prop (K tac) end;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2184
17510
5e3ce025e1a5 tuned simprocs;
wenzelm
parents: 17485
diff changeset
  2185
    val ss = get_simpset defs_thy;
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  2186
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2187
    fun sel_convs_prf () =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2188
      map (prove_simp false ss (sel_defs @ accessor_thms)) sel_conv_props;
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  2189
    val sel_convs = timeit_msg "record sel_convs proof:" sel_convs_prf;
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 34151
diff changeset
  2190
    fun sel_convs_standard_prf () = map Drule.export_without_context sel_convs;
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2191
    val sel_convs_standard =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2192
      timeit_msg "record sel_convs_standard proof:" sel_convs_standard_prf;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2193
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2194
    fun upd_convs_prf () =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2195
      map (prove_simp false ss (upd_defs @ updator_thms)) upd_conv_props;
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  2196
    val upd_convs = timeit_msg "record upd_convs proof:" upd_convs_prf;
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 34151
diff changeset
  2197
    fun upd_convs_standard_prf () = map Drule.export_without_context upd_convs;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2198
    val upd_convs_standard =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2199
      timeit_msg "record upd_convs_standard proof:" upd_convs_standard_prf;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2200
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2201
    fun get_upd_acc_congs () =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2202
      let
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
  2203
        val symdefs = map Thm.symmetric (sel_defs @ upd_defs);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2204
        val fold_ss = HOL_basic_ss addsimps symdefs;
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 34151
diff changeset
  2205
        val ua_congs = map (Drule.export_without_context o simplify fold_ss) upd_acc_cong_assists;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2206
      in (ua_congs RL [updacc_foldE], ua_congs RL [updacc_unfoldE]) end;
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2207
    val (fold_congs, unfold_congs) =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2208
      timeit_msg "record upd fold/unfold congs:" get_upd_acc_congs;
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  2209
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  2210
    val parent_induct = Option.map #induct_scheme (try List.last parents);
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2211
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2212
    fun induct_scheme_prf () =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2213
      prove_standard [] induct_scheme_prop
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2214
        (fn _ =>
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2215
          EVERY
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  2216
           [case parent_induct of NONE => all_tac | SOME ind => try_param_tac rN ind 1,
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2217
            try_param_tac rN ext_induct 1,
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2218
            asm_simp_tac HOL_basic_ss 1]);
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  2219
    val induct_scheme = timeit_msg "record induct_scheme proof:" induct_scheme_prf;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2220
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  2221
    fun induct_prf () =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2222
      let val (assm, concl) = induct_prop in
20248
7916ce5bb069 Goal.prove: more tactic arguments;
wenzelm
parents: 20071
diff changeset
  2223
        prove_standard [assm] concl (fn {prems, ...} =>
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2224
          try_param_tac rN induct_scheme 1
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26957
diff changeset
  2225
          THEN try_param_tac "more" @{thm unit.induct} 1
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2226
          THEN resolve_tac prems 1)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2227
      end;
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  2228
    val induct = timeit_msg "record induct proof:" induct_prf;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2229
33691
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  2230
    fun cases_scheme_prf () =
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2231
      let
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2232
        val _ $ (Pvar $ _) = concl_of induct_scheme;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2233
        val ind =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2234
          cterm_instantiate
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2235
            [(cterm_of defs_thy Pvar, cterm_of defs_thy
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2236
              (lambda w (HOLogic.imp $ HOLogic.mk_eq (r0, w) $ C)))]
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2237
            induct_scheme;
35625
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 35615
diff changeset
  2238
        in Object_Logic.rulify (mp OF [ind, refl]) end;
33691
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  2239
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  2240
    val cases_scheme_prf = future_forward_prf cases_scheme_prf cases_scheme_prop;
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  2241
    val cases_scheme = timeit_msg "record cases_scheme proof:" cases_scheme_prf;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2242
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  2243
    fun cases_prf () =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2244
      prove_standard [] cases_prop
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2245
        (fn _ =>
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2246
          try_param_tac rN cases_scheme 1 THEN
37136
e0c9d3e49e15 dropped legacy theorem bindings
haftmann
parents: 36960
diff changeset
  2247
          simp_all_tac HOL_basic_ss [@{thm unit_all_eq1}]);
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  2248
    val cases = timeit_msg "record cases proof:" cases_prf;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2249
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2250
    fun surjective_prf () =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2251
      let
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2252
        val leaf_ss = get_sel_upd_defs defs_thy addsimps (sel_defs @ (o_assoc :: id_o_apps));
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2253
        val init_ss = HOL_basic_ss addsimps ext_defs;
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2254
      in
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2255
        prove_standard [] surjective_prop
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  2256
          (fn _ =>
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2257
            EVERY
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2258
             [rtac surject_assist_idE 1,
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2259
              simp_tac init_ss 1,
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  2260
              REPEAT
34151
8d57ce46b3f7 prefer prefix "iso" over potentially misleading "is"; tuned
haftmann
parents: 33957
diff changeset
  2261
                (Iso_Tuple_Support.iso_tuple_intros_tac 1 ORELSE
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  2262
                  (rtac surject_assistI 1 THEN simp_tac leaf_ss 1))])
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2263
      end;
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2264
    val surjective = timeit_msg "record surjective proof:" surjective_prf;
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2265
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  2266
    fun split_meta_prf () =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2267
      prove false [] split_meta_prop
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  2268
        (fn _ =>
32975
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  2269
          EVERY1
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  2270
           [rtac equal_intr_rule, Goal.norm_hhf_tac,
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  2271
            etac meta_allE, atac,
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  2272
            rtac (prop_subst OF [surjective]),
84d105ad5adb simplified tactics;
wenzelm
parents: 32974
diff changeset
  2273
            REPEAT o etac meta_allE, atac]);
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  2274
    val split_meta = timeit_msg "record split_meta proof:" split_meta_prf;
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 34151
diff changeset
  2275
    fun split_meta_standardise () = Drule.export_without_context split_meta;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2276
    val split_meta_standard =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2277
      timeit_msg "record split_meta standard:" split_meta_standardise;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2278
33691
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  2279
    fun split_object_prf () =
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2280
      let
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2281
        val cPI= cterm_of defs_thy (lambda r0 (Trueprop (P $ r0)));
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2282
        val _ $ Abs (_, _, P $ _) = fst (Logic.dest_equals (concl_of split_meta_standard));
18858
ceb93f3af7f0 advanced translations: Context.generic;
wenzelm
parents: 18728
diff changeset
  2283
        val cP = cterm_of defs_thy P;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2284
        val split_meta' = cterm_instantiate [(cP, cPI)] split_meta_standard;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2285
        val (l, r) = HOLogic.dest_eq (HOLogic.dest_Trueprop split_object_prop);
18858
ceb93f3af7f0 advanced translations: Context.generic;
wenzelm
parents: 18728
diff changeset
  2286
        val cl = cterm_of defs_thy (HOLogic.mk_Trueprop l);
ceb93f3af7f0 advanced translations: Context.generic;
wenzelm
parents: 18728
diff changeset
  2287
        val cr = cterm_of defs_thy (HOLogic.mk_Trueprop r);
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2288
        val thl =
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
  2289
          Thm.assume cl                   (*All r. P r*) (* 1 *)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
  2290
          |> obj_to_meta_all              (*!!r. P r*)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
  2291
          |> Thm.equal_elim split_meta'   (*!!n m more. P (ext n m more)*)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
  2292
          |> meta_to_obj_all              (*All n m more. P (ext n m more)*) (* 2*)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
  2293
          |> Thm.implies_intr cl          (* 1 ==> 2 *)
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2294
        val thr =
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
  2295
          Thm.assume cr                                 (*All n m more. P (ext n m more)*)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
  2296
          |> obj_to_meta_all                            (*!!n m more. P (ext n m more)*)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
  2297
          |> Thm.equal_elim (Thm.symmetric split_meta') (*!!r. P r*)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
  2298
          |> meta_to_obj_all                            (*All r. P r*)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
  2299
          |> Thm.implies_intr cr                        (* 2 ==> 1 *)
33691
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  2300
     in thr COMP (thl COMP iffI) end;
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  2301
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  2302
3db22a5707f3 clarified quick-and-dirty usage in record package;
schirmer <schirmer@in.tum.de>
parents: 33612
diff changeset
  2303
    val split_object_prf = future_forward_prf split_object_prf split_object_prop;
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  2304
    val split_object = timeit_msg "record split_object proof:" split_object_prf;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2305
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  2306
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2307
    fun split_ex_prf () =
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2308
      let
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2309
        val ss = HOL_basic_ss addsimps [not_ex RS sym, Not_eq_iff];
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2310
        val P_nm = fst (dest_Free P);
32752
f65d74a264dd Initial response to feedback from Norbert, Makarius on record patch
tsewell@rubicon.NSW.bigpond.net.au
parents: 32749
diff changeset
  2311
        val not_P = cterm_of defs_thy (lambda r0 (HOLogic.mk_not (P $ r0)));
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2312
        val so' = named_cterm_instantiate ([(P_nm, not_P)]) split_object;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2313
        val so'' = simplify ss so';
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2314
      in
32799
7478ea535416 eliminated dead code, redundant bindings and parameters;
wenzelm
parents: 32770
diff changeset
  2315
        prove_standard [] split_ex_prop (fn _ => resolve_tac [so''] 1)
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2316
      end;
15012
28fa57b57209 Added reference record_definition_quick_and_dirty_sensitive, to
schirmer
parents: 14981
diff changeset
  2317
    val split_ex = timeit_msg "record split_ex proof:" split_ex_prf;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2318
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2319
    fun equality_tac thms =
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2320
      let
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2321
        val s' :: s :: eqs = rev thms;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2322
        val ss' = ss addsimps (s' :: s :: sel_convs_standard);
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2323
        val eqs' = map (simplify ss') eqs;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2324
      in simp_tac (HOL_basic_ss addsimps (s' :: s :: eqs')) 1 end;
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2325
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2326
    fun equality_prf () =
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2327
      prove_standard [] equality_prop (fn {context, ...} =>
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2328
        fn st =>
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2329
          let val [s, s'] = map #1 (rev (Tactic.innermost_params 1 st)) in
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2330
            st |> (res_inst_tac context [((rN, 0), s)] cases_scheme 1 THEN
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2331
              res_inst_tac context [((rN, 0), s')] cases_scheme 1 THEN
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2332
              Subgoal.FOCUS (fn {prems, ...} => equality_tac prems) context 1)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2333
             (*simp_all_tac ss (sel_convs) would also work but is less efficient*)
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2334
          end);
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2335
    val equality = timeit_msg "record equality proof:" equality_prf;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2336
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
  2337
    val ((([sel_convs', upd_convs', sel_defs', upd_defs',
32744
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
  2338
            fold_congs', unfold_congs',
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  2339
          splits' as [split_meta', split_object', split_ex'], derived_defs'],
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
  2340
          [surjective', equality']),
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
  2341
          [induct_scheme', induct', cases_scheme', cases']), thms_thy) =
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2342
      defs_thy
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39134
diff changeset
  2343
      |> (Global_Theory.add_thmss o map (Thm.no_attributes o apfst Binding.name))
15215
6bd87812537c tuned performance of record definition
schirmer
parents: 15203
diff changeset
  2344
         [("select_convs", sel_convs_standard),
32744
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
  2345
          ("update_convs", upd_convs_standard),
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2346
          ("select_defs", sel_defs),
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2347
          ("update_defs", upd_defs),
32744
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
  2348
          ("fold_congs", fold_congs),
50406c4951d9 Record patch imported and working.
Thomas Sewell <tsewell@nicta.com.au>
parents: 32743
diff changeset
  2349
          ("unfold_congs", unfold_congs),
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2350
          ("splits", [split_meta_standard, split_object, split_ex]),
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2351
          ("defs", derived_defs)]
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39134
diff changeset
  2352
      ||>> (Global_Theory.add_thms o map (Thm.no_attributes o apfst Binding.name))
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2353
          [("surjective", surjective),
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2354
           ("equality", equality)]
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39134
diff changeset
  2355
      ||>> (Global_Theory.add_thms o (map o apfst o apfst) Binding.name)
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2356
        [(("induct_scheme", induct_scheme), induct_type_global (suffix schemeN name)),
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2357
         (("induct", induct), induct_type_global name),
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2358
         (("cases_scheme", cases_scheme), cases_type_global (suffix schemeN name)),
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2359
         (("cases", cases), cases_type_global name)];
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2360
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2361
    val sel_upd_simps = sel_convs' @ upd_convs';
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2362
    val sel_upd_defs = sel_defs' @ upd_defs';
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2363
    val iffs = [ext_inject]
32743
c4e9a48bc50e Initial attempt at porting record update to repository Isabelle.
tsewell@rubicon.NSW.bigpond.net.au
parents: 32335
diff changeset
  2364
    val depth = parent_len + 1;
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  2365
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  2366
    val ([simps', iffs'], thms_thy') =
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2367
      thms_thy
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39134
diff changeset
  2368
      |> Global_Theory.add_thmss
33053
dabf9d1bb552 removed "nitpick_const_simp" attribute from Record's "simps";
blanchet
parents: 32952
diff changeset
  2369
          [((Binding.name "simps", sel_upd_simps), [Simplifier.simp_add]),
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  2370
           ((Binding.name "iffs", iffs), [iff_add])];
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  2371
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  2372
    val info =
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  2373
      make_info alphas parent fields extension
35138
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  2374
        ext_induct ext_inject ext_surjective ext_split ext_def
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  2375
        sel_convs' upd_convs' sel_defs' upd_defs' fold_congs' unfold_congs' splits' derived_defs'
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  2376
        surjective' equality' induct_scheme' induct' cases_scheme' cases' simps' iffs';
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  2377
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  2378
    val final_thy =
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  2379
      thms_thy'
ad213c602ec1 refined and exported record_info;
wenzelm
parents: 35137
diff changeset
  2380
      |> put_record name info
32764
690f9cccf232 replaced meta_iffD2 by existing Drule.equal_elim_rule2;
wenzelm
parents: 32763
diff changeset
  2381
      |> put_sel_upd names full_moreN depth sel_upd_simps sel_upd_defs (fold_congs', unfold_congs')
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  2382
      |> add_equalities extension_id equality'
15015
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
  2383
      |> add_extinjects ext_inject
c5768e8c4da4 * record_upd_simproc also simplifies trivial updates:
schirmer
parents: 15012
diff changeset
  2384
      |> add_extsplit extension_name ext_split
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  2385
      |> add_splits extension_id (split_meta', split_object', split_ex', induct_scheme')
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2386
      |> add_extfields extension_name (fields @ [(full_moreN, moreT)])
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2387
      |> add_fieldext (extension_name, snd extension) (names @ [full_moreN])
38533
8d23c7403699 use extension constant as formal constructor of logical record type
haftmann
parents: 38531
diff changeset
  2388
      |> add_code ext_tyco vs extT ext simps' ext_inject
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2389
      |> Sign.restore_naming thy;
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14643
diff changeset
  2390
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2391
  in final_thy end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2392
27278
129574589734 export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27239
diff changeset
  2393
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2394
(* add_record *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2395
36151
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2396
local
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2397
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2398
fun read_parent NONE ctxt = (NONE, ctxt)
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2399
  | read_parent (SOME raw_T) ctxt =
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2400
      (case ProofContext.read_typ_abbrev ctxt raw_T of
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2401
        Type (name, Ts) => (SOME (Ts, name), fold Variable.declare_typ Ts ctxt)
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2402
      | T => error ("Bad parent record specification: " ^ Syntax.string_of_typ ctxt T));
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2403
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2404
fun prep_field prep (x, T, mx) = (x, prep T, mx)
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2405
  handle ERROR msg =>
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2406
    cat_error msg ("The error(s) above occurred in record field " ^ quote (Binding.str_of x));
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2407
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2408
fun read_field raw_field ctxt =
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2409
  let val field as (_, T, _) = prep_field (Syntax.read_typ ctxt) raw_field
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2410
  in (field, Variable.declare_typ T ctxt) end;
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2411
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2412
in
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2413
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2414
fun add_record quiet_mode (params, binding) raw_parent raw_fields thy =
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2415
  let
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2416
    val _ = Theory.requires thy "Record" "record definitions";
35136
34206672b168 modernized signature -- proper binding;
wenzelm
parents: 35135
diff changeset
  2417
    val _ =
34206672b168 modernized signature -- proper binding;
wenzelm
parents: 35135
diff changeset
  2418
      if quiet_mode then ()
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2419
      else writeln ("Defining record " ^ quote (Binding.str_of binding) ^ " ...");
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2420
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36173
diff changeset
  2421
    val ctxt = ProofContext.init_global thy;
36151
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2422
    fun cert_typ T = Type.no_tvars (ProofContext.cert_typ ctxt T)
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2423
      handle TYPE (msg, _, _) => error msg;
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2424
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2425
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2426
    (* specification *)
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2427
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2428
    val parent = Option.map (apfst (map cert_typ)) raw_parent
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2429
      handle ERROR msg =>
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2430
        cat_error msg ("The error(s) above occurred in parent record specification");
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2431
    val parent_args = (case parent of SOME (Ts, _) => Ts | NONE => []);
12247
9b029789aff6 derive cases/induct rules for ``more'' parts;
wenzelm
parents: 12129
diff changeset
  2432
    val parents = add_parents thy parent [];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2433
36151
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2434
    val bfields = map (prep_field cert_typ) raw_fields;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2435
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2436
    (* errors *)
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2437
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2438
    val name = Sign.full_name thy binding;
17261
193b84a70ca4 curried_lookup/update;
wenzelm
parents: 17257
diff changeset
  2439
    val err_dup_record =
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  2440
      if is_none (get_info thy name) then []
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
  2441
      else ["Duplicate definition of record " ^ quote name];
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
  2442
36151
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2443
    val spec_frees = fold Term.add_tfreesT (parent_args @ map #2 bfields) [];
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2444
    val err_extra_frees =
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2445
      (case subtract (op =) params spec_frees of
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2446
        [] => []
36151
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2447
      | extras => ["Extra free type variable(s) " ^
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2448
          commas (map (Syntax.string_of_typ ctxt o TFree) extras)]);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2449
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
  2450
    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
  2451
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2452
    val err_dup_fields =
35136
34206672b168 modernized signature -- proper binding;
wenzelm
parents: 35135
diff changeset
  2453
      (case duplicates Binding.eq_name (map #1 bfields) of
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2454
        [] => []
35136
34206672b168 modernized signature -- proper binding;
wenzelm
parents: 35135
diff changeset
  2455
      | dups => ["Duplicate field(s) " ^ commas_quote (map Binding.str_of dups)]);
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
  2456
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
  2457
    val err_bad_fields =
35136
34206672b168 modernized signature -- proper binding;
wenzelm
parents: 35135
diff changeset
  2458
      if forall (not_equal moreN o Binding.name_of o #1) bfields then []
4890
f0a24bad990a concrete syntax for record terms;
wenzelm
parents: 4867
diff changeset
  2459
      else ["Illegal field name " ^ quote moreN];
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2460
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2461
    val errs =
36151
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2462
      err_dup_record @ err_extra_frees @ err_no_fields @ err_dup_fields @ err_bad_fields;
32761
54fee94b2b29 tuned whitespace -- recover basic Isabelle conventions;
wenzelm
parents: 32760
diff changeset
  2463
    val _ = if null errs then () else error (cat_lines errs);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2464
  in
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  2465
    thy |> definition (params, binding) parent parents bfields
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2466
  end
35239
0dfec017bc83 authentic term syntax;
wenzelm
parents: 35232
diff changeset
  2467
  handle ERROR msg => cat_error msg ("Failed to define record " ^ quote (Binding.str_of binding));
35136
34206672b168 modernized signature -- proper binding;
wenzelm
parents: 35135
diff changeset
  2468
36153
1ac501e16a6a replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
wenzelm
parents: 36151
diff changeset
  2469
fun add_record_cmd quiet_mode (raw_params, binding) raw_parent raw_fields thy =
36151
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2470
  let
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36173
diff changeset
  2471
    val ctxt = ProofContext.init_global thy;
36153
1ac501e16a6a replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
wenzelm
parents: 36151
diff changeset
  2472
    val params = map (apsnd (Typedecl.read_constraint ctxt)) raw_params;
1ac501e16a6a replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
wenzelm
parents: 36151
diff changeset
  2473
    val ctxt1 = fold (Variable.declare_typ o TFree) params ctxt;
1ac501e16a6a replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
wenzelm
parents: 36151
diff changeset
  2474
    val (parent, ctxt2) = read_parent raw_parent ctxt1;
1ac501e16a6a replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
wenzelm
parents: 36151
diff changeset
  2475
    val (fields, ctxt3) = fold_map read_field raw_fields ctxt2;
1ac501e16a6a replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
wenzelm
parents: 36151
diff changeset
  2476
    val params' = map (ProofContext.check_tfree ctxt3) params;
36151
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2477
  in thy |> add_record quiet_mode (params', binding) parent fields end;
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2478
b89a2a05a3ce modernized treatment of sort constraints in specification;
wenzelm
parents: 36137
diff changeset
  2479
end;
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2480
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
  2481
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  2482
(* setup theory *)
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2483
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2484
val setup =
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24624
diff changeset
  2485
  Sign.add_trfuns ([], parse_translation, [], []) #>
35146
f7bf73b0d7e5 simplified/clarified record translations;
wenzelm
parents: 35145
diff changeset
  2486
  Sign.add_advanced_trfuns ([], advanced_parse_translation, [], []) #>
26496
49ae9456eba9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26477
diff changeset
  2487
  Simplifier.map_simpset (fn ss =>
38012
3ca193a6ae5a delete structure Basic_Record; avoid `record` in names in structure Record
haftmann
parents: 37781
diff changeset
  2488
    ss addsimprocs [simproc, upd_simproc, eq_simproc]);
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2489
32335
41fe1c93f218 tuned spacing of sections;
wenzelm
parents: 32283
diff changeset
  2490
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  2491
(* outer syntax *)
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  2492
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24830
diff changeset
  2493
val _ =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36945
diff changeset
  2494
  Outer_Syntax.command "record" "define extensible record" Keyword.thy_decl
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36945
diff changeset
  2495
    (Parse.type_args_constrained -- Parse.binding --
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36945
diff changeset
  2496
      (Parse.$$$ "=" |-- Scan.option (Parse.typ --| Parse.$$$ "+") --
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36945
diff changeset
  2497
        Scan.repeat1 Parse.const_binding)
35136
34206672b168 modernized signature -- proper binding;
wenzelm
parents: 35135
diff changeset
  2498
    >> (fn (x, (y, z)) => Toplevel.theory (add_record_cmd false x y z)));
6358
92dbe243555f outer syntax for 'record';
wenzelm
parents: 6092
diff changeset
  2499
4867
9be2bf0ce909 package extensible records with structural subtyping in HOL -- still
wenzelm
parents:
diff changeset
  2500
end;