src/Pure/Tools/codegen_consts.ML
author haftmann
Thu, 21 Dec 2006 13:55:15 +0100
changeset 21895 6cbc0f69a21c
parent 21463 42dd50268c8b
child 22032 979671292fbe
permissions -rw-r--r--
clarified code
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
20387
haftmann
parents:
diff changeset
     1
(*  Title:      Pure/Tools/codegen_consts.ML
haftmann
parents:
diff changeset
     2
    ID:         $Id$
haftmann
parents:
diff changeset
     3
    Author:     Florian Haftmann, TU Muenchen
haftmann
parents:
diff changeset
     4
20855
9f60d493c8fe clarified header comments
haftmann
parents: 20704
diff changeset
     5
Identifying constants by name plus normalized type instantantiation schemes.
9f60d493c8fe clarified header comments
haftmann
parents: 20704
diff changeset
     6
Type normalization is compatible with overloading discipline and user-defined
9f60d493c8fe clarified header comments
haftmann
parents: 20704
diff changeset
     7
ad-hoc overloading. Convenient data structures for constants.
20387
haftmann
parents:
diff changeset
     8
*)
haftmann
parents:
diff changeset
     9
haftmann
parents:
diff changeset
    10
signature CODEGEN_CONSTS =
haftmann
parents:
diff changeset
    11
sig
haftmann
parents:
diff changeset
    12
  type const = string * typ list (*type instantiations*)
haftmann
parents:
diff changeset
    13
  val const_ord: const * const -> order
haftmann
parents:
diff changeset
    14
  val eq_const: const * const -> bool
haftmann
parents:
diff changeset
    15
  structure Consttab: TABLE
20600
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    16
  val inst_of_typ: theory -> string * typ -> const
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    17
  val typ_of_inst: theory -> const -> string * typ
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    18
  val norm: theory -> const -> const
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    19
  val norm_of_typ: theory -> string * typ -> const
20389
8b6ecb22ef35 cleanup
haftmann
parents: 20387
diff changeset
    20
  val find_def: theory -> const
20600
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    21
    -> ((string (*theory name*) * thm) * typ list) option
21463
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21118
diff changeset
    22
  val instance_dict: theory -> class * string
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21118
diff changeset
    23
    -> (string * sort) list * (string * typ) list
20704
a56f0743b3ee cleaned up
haftmann
parents: 20600
diff changeset
    24
  val disc_typ_of_classop: theory -> const -> typ
a56f0743b3ee cleaned up
haftmann
parents: 20600
diff changeset
    25
  val disc_typ_of_const: theory -> (const -> typ) -> const -> typ
21118
d9789a87825d cleanup
haftmann
parents: 20855
diff changeset
    26
  val consts_of: theory -> term -> const list
20387
haftmann
parents:
diff changeset
    27
  val read_const: theory -> string -> const
haftmann
parents:
diff changeset
    28
  val string_of_const: theory -> const -> string
20600
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    29
  val raw_string_of_const: const -> string
20387
haftmann
parents:
diff changeset
    30
  val string_of_const_typ: theory -> string * typ -> string
haftmann
parents:
diff changeset
    31
end;
haftmann
parents:
diff changeset
    32
haftmann
parents:
diff changeset
    33
structure CodegenConsts: CODEGEN_CONSTS =
haftmann
parents:
diff changeset
    34
struct
haftmann
parents:
diff changeset
    35
haftmann
parents:
diff changeset
    36
haftmann
parents:
diff changeset
    37
(* basic data structures *)
haftmann
parents:
diff changeset
    38
haftmann
parents:
diff changeset
    39
type const = string * typ list (*type instantiations*);
haftmann
parents:
diff changeset
    40
val const_ord = prod_ord fast_string_ord (list_ord Term.typ_ord);
haftmann
parents:
diff changeset
    41
val eq_const = is_equal o const_ord;
haftmann
parents:
diff changeset
    42
haftmann
parents:
diff changeset
    43
structure Consttab =
haftmann
parents:
diff changeset
    44
  TableFun(
haftmann
parents:
diff changeset
    45
    type key = string * typ list;
haftmann
parents:
diff changeset
    46
    val ord = const_ord;
haftmann
parents:
diff changeset
    47
  );
haftmann
parents:
diff changeset
    48
haftmann
parents:
diff changeset
    49
21463
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21118
diff changeset
    50
(* type instantiations, overloading, dictionary values *)
20387
haftmann
parents:
diff changeset
    51
20600
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    52
fun inst_of_typ thy (c_ty as (c, ty)) =
20387
haftmann
parents:
diff changeset
    53
  (c, Consts.typargs (Sign.consts_of thy) c_ty);
haftmann
parents:
diff changeset
    54
20600
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    55
fun typ_of_inst thy (c_tys as (c, tys)) =
20387
haftmann
parents:
diff changeset
    56
  (c, Consts.instance (Sign.consts_of thy) c_tys);
haftmann
parents:
diff changeset
    57
20389
8b6ecb22ef35 cleanup
haftmann
parents: 20387
diff changeset
    58
fun find_def thy (c, tys) =
20387
haftmann
parents:
diff changeset
    59
  let
haftmann
parents:
diff changeset
    60
    val specs = Defs.specifications_of (Theory.defs_of thy) c;
20389
8b6ecb22ef35 cleanup
haftmann
parents: 20387
diff changeset
    61
    val typ_instance = case AxClass.class_of_param thy c
8b6ecb22ef35 cleanup
haftmann
parents: 20387
diff changeset
    62
     of SOME _ => let
8b6ecb22ef35 cleanup
haftmann
parents: 20387
diff changeset
    63
          fun instance_tycos (Type (tyco1, _), Type (tyco2, _)) = tyco1 = tyco2
8b6ecb22ef35 cleanup
haftmann
parents: 20387
diff changeset
    64
            | instance_tycos (_ , TVar _) = true
8b6ecb22ef35 cleanup
haftmann
parents: 20387
diff changeset
    65
            | instance_tycos ty_ty = Sign.typ_instance thy ty_ty;
8b6ecb22ef35 cleanup
haftmann
parents: 20387
diff changeset
    66
        in instance_tycos end
20600
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    67
      | NONE =>  Sign.typ_instance thy;
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    68
    fun get_def (_, { is_def, thyname, name, lhs, rhs }) =
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    69
      if is_def andalso forall typ_instance (tys ~~ lhs) then
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    70
        case try (Thm.get_axiom_i thy) name
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    71
         of SOME thm => SOME ((thyname, thm), lhs)
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    72
          | NONE => NONE
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    73
      else NONE
20387
haftmann
parents:
diff changeset
    74
  in
20600
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    75
    get_first get_def specs
20387
haftmann
parents:
diff changeset
    76
  end;
haftmann
parents:
diff changeset
    77
20389
8b6ecb22ef35 cleanup
haftmann
parents: 20387
diff changeset
    78
fun mk_classop_typinst thy (c, tyco) =
8b6ecb22ef35 cleanup
haftmann
parents: 20387
diff changeset
    79
  (c, [Type (tyco, map (fn v => TVar ((v, 0), Sign.defaultS thy (*for monotonicity*)))
8b6ecb22ef35 cleanup
haftmann
parents: 20387
diff changeset
    80
    (Name.invents Name.context "'a" (Sign.arity_number thy tyco)))]);
8b6ecb22ef35 cleanup
haftmann
parents: 20387
diff changeset
    81
20600
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    82
fun norm thy (c, insts) =
20387
haftmann
parents:
diff changeset
    83
  let
haftmann
parents:
diff changeset
    84
    fun disciplined _ [(Type (tyco, _))] =
20389
8b6ecb22ef35 cleanup
haftmann
parents: 20387
diff changeset
    85
          mk_classop_typinst thy (c, tyco)
20387
haftmann
parents:
diff changeset
    86
      | disciplined sort _ =
20389
8b6ecb22ef35 cleanup
haftmann
parents: 20387
diff changeset
    87
          (c, [TVar (("'a", 0), sort)]);
20387
haftmann
parents:
diff changeset
    88
    fun ad_hoc c tys =
20389
8b6ecb22ef35 cleanup
haftmann
parents: 20387
diff changeset
    89
      case find_def thy (c, tys)
20387
haftmann
parents:
diff changeset
    90
       of SOME (_, tys) => (c, tys)
20600
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    91
        | NONE => inst_of_typ thy (c, Sign.the_const_type thy c);
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    92
  in case AxClass.class_of_param thy c
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    93
     of SOME class => disciplined [class] insts
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    94
      | _ => ad_hoc c insts
20387
haftmann
parents:
diff changeset
    95
  end;
haftmann
parents:
diff changeset
    96
20600
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    97
fun norm_of_typ thy (c, ty) =
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    98
  norm thy (c, Consts.typargs (Sign.consts_of thy) (c, ty));
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
    99
21463
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21118
diff changeset
   100
fun instance_dict thy (class, tyco) =
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21118
diff changeset
   101
  let
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21118
diff changeset
   102
    val (var, cs) = AxClass.params_of_class thy class;
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21118
diff changeset
   103
    val sort_args = Name.names (Name.declare var Name.context) "'a"
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21118
diff changeset
   104
      (Sign.arity_sorts thy tyco [class]);
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21118
diff changeset
   105
    val ty_inst = Type (tyco, map TFree sort_args);
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21118
diff changeset
   106
    val inst_signs = (map o apsnd o map_type_tfree) (K ty_inst) cs;
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21118
diff changeset
   107
  in (sort_args, inst_signs) end;
42dd50268c8b completed class parameter handling in axclass.ML
haftmann
parents: 21118
diff changeset
   108
21895
6cbc0f69a21c clarified code
haftmann
parents: 21463
diff changeset
   109
fun disc_typ_of_classop thy (c, [ty]) = 
6cbc0f69a21c clarified code
haftmann
parents: 21463
diff changeset
   110
  let
6cbc0f69a21c clarified code
haftmann
parents: 21463
diff changeset
   111
    val class = (the o AxClass.class_of_param thy) c;
6cbc0f69a21c clarified code
haftmann
parents: 21463
diff changeset
   112
    val cs = case ty
6cbc0f69a21c clarified code
haftmann
parents: 21463
diff changeset
   113
     of TVar _ => snd (AxClass.params_of_class thy class)
6cbc0f69a21c clarified code
haftmann
parents: 21463
diff changeset
   114
      | TFree _ => snd (AxClass.params_of_class thy class)
6cbc0f69a21c clarified code
haftmann
parents: 21463
diff changeset
   115
      | Type (tyco, _) => snd (instance_dict thy (class, tyco));
6cbc0f69a21c clarified code
haftmann
parents: 21463
diff changeset
   116
  in
6cbc0f69a21c clarified code
haftmann
parents: 21463
diff changeset
   117
    (Logic.varifyT o the o AList.lookup (op =) cs) c
6cbc0f69a21c clarified code
haftmann
parents: 21463
diff changeset
   118
  end;
20387
haftmann
parents:
diff changeset
   119
20704
a56f0743b3ee cleaned up
haftmann
parents: 20600
diff changeset
   120
fun disc_typ_of_const thy f (const as (c, [ty])) =
a56f0743b3ee cleaned up
haftmann
parents: 20600
diff changeset
   121
      if (is_some o AxClass.class_of_param thy) c
a56f0743b3ee cleaned up
haftmann
parents: 20600
diff changeset
   122
      then disc_typ_of_classop thy const
a56f0743b3ee cleaned up
haftmann
parents: 20600
diff changeset
   123
      else f const
a56f0743b3ee cleaned up
haftmann
parents: 20600
diff changeset
   124
  | disc_typ_of_const thy f const = f const;
a56f0743b3ee cleaned up
haftmann
parents: 20600
diff changeset
   125
a56f0743b3ee cleaned up
haftmann
parents: 20600
diff changeset
   126
fun consts_of thy t =
21118
d9789a87825d cleanup
haftmann
parents: 20855
diff changeset
   127
  fold_aterms (fn Const c => cons (norm_of_typ thy c) | _ => I) t []
20704
a56f0743b3ee cleaned up
haftmann
parents: 20600
diff changeset
   128
20387
haftmann
parents:
diff changeset
   129
haftmann
parents:
diff changeset
   130
(* reading constants as terms *)
haftmann
parents:
diff changeset
   131
haftmann
parents:
diff changeset
   132
fun read_const_typ thy raw_t =
haftmann
parents:
diff changeset
   133
  let
haftmann
parents:
diff changeset
   134
    val t = Sign.read_term thy raw_t
haftmann
parents:
diff changeset
   135
  in case try dest_Const t
20600
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
   136
   of SOME c_ty => (typ_of_inst thy o norm_of_typ thy) c_ty
20389
8b6ecb22ef35 cleanup
haftmann
parents: 20387
diff changeset
   137
    | NONE => error ("Not a constant: " ^ Sign.string_of_term thy t)
20387
haftmann
parents:
diff changeset
   138
  end;
haftmann
parents:
diff changeset
   139
haftmann
parents:
diff changeset
   140
fun read_const thy =
20600
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
   141
  norm_of_typ thy o read_const_typ thy;
20387
haftmann
parents:
diff changeset
   142
haftmann
parents:
diff changeset
   143
haftmann
parents:
diff changeset
   144
(* printing constants *)
haftmann
parents:
diff changeset
   145
haftmann
parents:
diff changeset
   146
fun string_of_const thy (c, tys) =
haftmann
parents:
diff changeset
   147
  space_implode " " (Sign.extern_const thy c
haftmann
parents:
diff changeset
   148
    :: map (enclose "[" "]" o Sign.string_of_typ thy) tys);
haftmann
parents:
diff changeset
   149
20600
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
   150
fun raw_string_of_const (c, tys) =
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
   151
  space_implode " " (c
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
   152
    :: map (enclose "[" "]" o Display.raw_string_of_typ) tys);
6d75e02ed285 added codegen_data
haftmann
parents: 20456
diff changeset
   153
20387
haftmann
parents:
diff changeset
   154
fun string_of_const_typ thy (c, ty) =
haftmann
parents:
diff changeset
   155
  string_of_const thy (c, Consts.typargs (Sign.consts_of thy) (c, ty));
haftmann
parents:
diff changeset
   156
haftmann
parents:
diff changeset
   157
end;