src/HOL/Tools/typecopy_package.ML
author haftmann
Wed, 05 Dec 2007 14:15:45 +0100
changeset 25534 d0b74fdd6067
parent 25506 c9bea6426932
child 25569 c597835d5de4
permissions -rw-r--r--
simplified infrastructure for code generator operational equality
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Tools/typecopy_package.ML
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
     2
    ID:         $Id$
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
     3
    Author:     Florian Haftmann, TU Muenchen
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
     4
20855
9f60d493c8fe clarified header comments
haftmann
parents: 20835
diff changeset
     5
Introducing copies of types using trivial typedefs; datatype-like abstraction.
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
     6
*)
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
     7
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
     8
signature TYPECOPY_PACKAGE =
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
     9
sig
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    10
  type info = {
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    11
    vs: (string * sort) list,
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    12
    constr: string,
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    13
    typ: typ,
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    14
    inject: thm,
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    15
    proj: string * typ,
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    16
    proj_def: thm
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    17
  }
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    18
  val add_typecopy: bstring * string list -> typ -> (bstring * bstring) option
20483
04aa552a83bc TypedefPackage.add_typedef_* now yields name of introduced type constructor
haftmann
parents: 20426
diff changeset
    19
    -> theory -> (string * info) * theory
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    20
  val get_typecopies: theory -> string list
25506
haftmann
parents: 25489
diff changeset
    21
  val get_info: theory -> string -> info option
25489
5b0625f6e324 simplified interpretations
haftmann
parents: 24711
diff changeset
    22
  val interpretation: (string -> theory -> theory) -> theory -> theory
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    23
  val print_typecopies: theory -> unit
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    24
  val setup: theory -> theory
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    25
end;
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    26
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    27
structure TypecopyPackage: TYPECOPY_PACKAGE =
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    28
struct
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    29
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    30
(* theory data *)
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    31
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    32
type info = {
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    33
  vs: (string * sort) list,
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    34
  constr: string,
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    35
  typ: typ,
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    36
  inject: thm,
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    37
  proj: string * typ,
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    38
  proj_def: thm
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    39
};
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    40
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    41
structure TypecopyData = TheoryDataFun
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    42
(
24626
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24624
diff changeset
    43
  type T = info Symtab.table;
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24624
diff changeset
    44
  val empty = Symtab.empty;
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    45
  val copy = I;
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    46
  val extend = I;
24626
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24624
diff changeset
    47
  fun merge _ = Symtab.merge (K true);
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    48
);
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    49
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    50
fun print_typecopies thy =
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    51
  let
24626
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24624
diff changeset
    52
    val tab = TypecopyData.get thy;
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    53
    fun mk (tyco, { vs, constr, typ, proj = (proj, _), ... } : info) =
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    54
      (Pretty.block o Pretty.breaks) [
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    55
        Sign.pretty_typ thy (Type (tyco, map TFree vs)),
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    56
        Pretty.str "=",
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    57
        (Pretty.str o Sign.extern_const thy) constr,
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    58
        Sign.pretty_typ thy typ,
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    59
        Pretty.block [Pretty.str "(", (Pretty.str o Sign.extern_const thy) proj, Pretty.str  ")"]];
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    60
    in
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    61
      (Pretty.writeln o Pretty.block o Pretty.fbreaks)
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
    62
        (Pretty.str "type copies:" :: map mk (Symtab.dest tab))
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    63
    end;
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    64
24626
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24624
diff changeset
    65
val get_typecopies = Symtab.keys o TypecopyData.get;
25506
haftmann
parents: 25489
diff changeset
    66
val get_info = Symtab.lookup o TypecopyData.get;
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    67
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    68
25489
5b0625f6e324 simplified interpretations
haftmann
parents: 24711
diff changeset
    69
(* interpretation of type copies *)
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    70
24711
e8bba7723858 simplified interpretation setup;
wenzelm
parents: 24626
diff changeset
    71
structure TypecopyInterpretation = InterpretationFun(type T = string val eq = op =);
25489
5b0625f6e324 simplified interpretations
haftmann
parents: 24711
diff changeset
    72
val interpretation = TypecopyInterpretation.interpretation;
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    73
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    74
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    75
(* add a type copy *)
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    76
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    77
local
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    78
20483
04aa552a83bc TypedefPackage.add_typedef_* now yields name of introduced type constructor
haftmann
parents: 20426
diff changeset
    79
fun gen_add_typecopy prep_typ (raw_tyco, raw_vs) raw_ty constr_proj thy =
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    80
  let
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    81
    val ty = prep_typ thy raw_ty;
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    82
    val vs = AList.make (the_default HOLogic.typeS o AList.lookup (op =) (typ_tfrees ty)) raw_vs;
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    83
    val tac = Tactic.rtac UNIV_witness 1;
20483
04aa552a83bc TypedefPackage.add_typedef_* now yields name of introduced type constructor
haftmann
parents: 20426
diff changeset
    84
    fun add_info tyco ( { abs_type = ty_abs, rep_type = ty_rep, Abs_name = c_abs,
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    85
      Rep_name = c_rep, Abs_inject = inject,
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21675
diff changeset
    86
      Abs_inverse = inverse, ... } : TypedefPackage.info ) thy =
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    87
        let
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    88
          val exists_thm =
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    89
            UNIV_I
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    90
            |> Drule.instantiate' [SOME (ctyp_of thy (Logic.varifyT ty_rep))] [];
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    91
          val inject' = inject OF [exists_thm, exists_thm];
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    92
          val proj_def = inverse OF [exists_thm];
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    93
          val info = {
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    94
            vs = vs,
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    95
            constr = c_abs,
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    96
            typ = ty_rep,
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    97
            inject = inject',
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    98
            proj = (c_rep, ty_abs --> ty_rep),
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
    99
            proj_def = proj_def
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
   100
          };
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
   101
        in
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
   102
          thy
24626
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24624
diff changeset
   103
          |> (TypecopyData.map o Symtab.update_new) (tyco, info)
24711
e8bba7723858 simplified interpretation setup;
wenzelm
parents: 24626
diff changeset
   104
          |> TypecopyInterpretation.data tyco
20483
04aa552a83bc TypedefPackage.add_typedef_* now yields name of introduced type constructor
haftmann
parents: 20426
diff changeset
   105
          |> pair (tyco, info)
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
   106
        end
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
   107
  in
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
   108
    thy
20483
04aa552a83bc TypedefPackage.add_typedef_* now yields name of introduced type constructor
haftmann
parents: 20426
diff changeset
   109
    |> setmp TypedefPackage.quiet_mode true
04aa552a83bc TypedefPackage.add_typedef_* now yields name of introduced type constructor
haftmann
parents: 20426
diff changeset
   110
        (TypedefPackage.add_typedef_i false (SOME raw_tyco) (raw_tyco, map fst vs, NoSyn)
04aa552a83bc TypedefPackage.add_typedef_* now yields name of introduced type constructor
haftmann
parents: 20426
diff changeset
   111
          (HOLogic.mk_UNIV ty) (Option.map swap constr_proj)) tac
04aa552a83bc TypedefPackage.add_typedef_* now yields name of introduced type constructor
haftmann
parents: 20426
diff changeset
   112
    |-> (fn (tyco, info) => add_info tyco info)
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
   113
  end;
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
   114
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
   115
in
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
   116
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
   117
val add_typecopy = gen_add_typecopy Sign.certify_typ;
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
   118
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
   119
end;
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
   120
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
   121
25489
5b0625f6e324 simplified interpretations
haftmann
parents: 24711
diff changeset
   122
(* code generator setup *)
20835
haftmann
parents: 20597
diff changeset
   123
25489
5b0625f6e324 simplified interpretations
haftmann
parents: 24711
diff changeset
   124
fun add_typecopy_spec tyco thy =
5b0625f6e324 simplified interpretations
haftmann
parents: 24711
diff changeset
   125
  let
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25506
diff changeset
   126
    val SOME { constr, proj_def, inject, vs, ... } = get_info thy tyco;
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25506
diff changeset
   127
    val sorts_eq =
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25506
diff changeset
   128
      map (curry (Sorts.inter_sort (Sign.classes_of thy)) [HOLogic.class_eq] o snd) vs;
25506
haftmann
parents: 25489
diff changeset
   129
    val ty = Logic.unvarifyT (Sign.the_const_type thy constr);
25489
5b0625f6e324 simplified interpretations
haftmann
parents: 24711
diff changeset
   130
  in
25506
haftmann
parents: 25489
diff changeset
   131
    thy
haftmann
parents: 25489
diff changeset
   132
    |> Code.add_datatype [(constr, ty)]
haftmann
parents: 25489
diff changeset
   133
    |> Code.add_func proj_def
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25506
diff changeset
   134
    |> Instance.instantiate ([tyco], sorts_eq, [HOLogic.class_eq]) (pair ())
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25506
diff changeset
   135
         ((K o K) (Class.intro_classes_tac []))
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25506
diff changeset
   136
    |> Code.add_func (CodeUnit.constrain_thm [HOLogic.class_eq] inject)
25489
5b0625f6e324 simplified interpretations
haftmann
parents: 24711
diff changeset
   137
  end;
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
   138
25489
5b0625f6e324 simplified interpretations
haftmann
parents: 24711
diff changeset
   139
val setup =
5b0625f6e324 simplified interpretations
haftmann
parents: 24711
diff changeset
   140
  TypecopyInterpretation.init
5b0625f6e324 simplified interpretations
haftmann
parents: 24711
diff changeset
   141
  #> interpretation add_typecopy_spec
20426
9ffea7a8b31c added typecopy_package
haftmann
parents:
diff changeset
   142
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22484
diff changeset
   143
end;