src/Tools/code/code_funcgr.ML
author haftmann
Wed, 15 Aug 2007 08:57:42 +0200
changeset 24283 8ca96f4e49cd
parent 24219 e558fe311376
child 24423 ae9cd0e92423
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     1
(*  Title:      Tools/code/code_funcgr.ML
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     2
    ID:         $Id$
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     3
    Author:     Florian Haftmann, TU Muenchen
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     4
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     5
Retrieving, normalizing and structuring defining equations in graph
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     6
with explicit dependencies.
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     7
*)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     8
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
     9
signature CODE_FUNCGR =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    10
sig
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    11
  type T
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    12
  val timing: bool ref
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    13
  val funcs: T -> CodeUnit.const -> thm list
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    14
  val typ: T -> CodeUnit.const -> typ
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    15
  val all: T -> CodeUnit.const list
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    16
  val pretty: theory -> T -> Pretty.T
24283
haftmann
parents: 24219
diff changeset
    17
  val make: theory -> CodeUnit.const list -> T
haftmann
parents: 24219
diff changeset
    18
  val make_consts: theory -> CodeUnit.const list -> CodeUnit.const list * T
haftmann
parents: 24219
diff changeset
    19
  val eval_conv: theory -> (T -> cterm -> thm) -> cterm -> thm
haftmann
parents: 24219
diff changeset
    20
  val eval_term: theory -> (T -> cterm -> 'a) -> cterm -> 'a
haftmann
parents: 24219
diff changeset
    21
  val intervene: theory -> T -> T
haftmann
parents: 24219
diff changeset
    22
    (*FIXME drop intervene as soon as possible*)
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    23
  structure Constgraph : GRAPH
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    24
end
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    25
24283
haftmann
parents: 24219
diff changeset
    26
structure CodeFuncgr : CODE_FUNCGR =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    27
struct
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    28
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    29
(** the graph type **)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    30
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    31
structure Constgraph = GraphFun (
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    32
  type key = CodeUnit.const;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    33
  val ord = CodeUnit.const_ord;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    34
);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    35
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    36
type T = (typ * thm list) Constgraph.T;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    37
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    38
fun funcs funcgr =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    39
  these o Option.map snd o try (Constgraph.get_node funcgr);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    40
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    41
fun typ funcgr =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    42
  fst o Constgraph.get_node funcgr;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    43
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    44
fun all funcgr = Constgraph.keys funcgr;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    45
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    46
fun pretty thy funcgr =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    47
  AList.make (snd o Constgraph.get_node funcgr) (Constgraph.keys funcgr)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    48
  |> (map o apfst) (CodeUnit.string_of_const thy)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    49
  |> sort (string_ord o pairself fst)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    50
  |> map (fn (s, thms) =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    51
       (Pretty.block o Pretty.fbreaks) (
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    52
         Pretty.str s
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    53
         :: map Display.pretty_thm thms
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    54
       ))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    55
  |> Pretty.chunks;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    56
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    57
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    58
(** generic combinators **)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    59
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    60
fun fold_consts f thms =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    61
  thms
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    62
  |> maps (op :: o swap o apfst (snd o strip_comb) o Logic.dest_equals o Thm.plain_prop_of)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    63
  |> (fold o fold_aterms) (fn Const c => f c | _ => I);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    64
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    65
fun consts_of (const, []) = []
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    66
  | consts_of (const, thms as thm :: _) = 
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    67
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    68
        val thy = Thm.theory_of_thm thm;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    69
        val is_refl = curry CodeUnit.eq_const const;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    70
        fun the_const c = case try (CodeUnit.const_of_cexpr thy) c
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    71
         of SOME const => if is_refl const then I else insert CodeUnit.eq_const const
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    72
          | NONE => I
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    73
      in fold_consts the_const thms [] end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    74
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    75
fun insts_of thy algebra c ty_decl ty =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    76
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    77
    val tys_decl = Sign.const_typargs thy (c, ty_decl);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    78
    val tys = Sign.const_typargs thy (c, ty);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    79
    fun class_relation (x, _) _ = x;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    80
    fun type_constructor tyco xs class =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    81
      (tyco, class) :: maps (maps fst) xs;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    82
    fun type_variable (TVar (_, sort)) = map (pair []) sort
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    83
      | type_variable (TFree (_, sort)) = map (pair []) sort;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    84
    fun mk_inst ty (TVar (_, sort)) = cons (ty, sort)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    85
      | mk_inst ty (TFree (_, sort)) = cons (ty, sort)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    86
      | mk_inst (Type (_, tys1)) (Type (_, tys2)) = fold2 mk_inst tys1 tys2;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    87
    fun of_sort_deriv (ty, sort) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    88
      Sorts.of_sort_derivation (Sign.pp thy) algebra
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    89
        { class_relation = class_relation, type_constructor = type_constructor,
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    90
          type_variable = type_variable }
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    91
        (ty, sort)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    92
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    93
    flat (maps of_sort_deriv (fold2 mk_inst tys tys_decl []))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    94
  end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    95
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    96
fun drop_classes thy tfrees thm =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    97
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    98
    val (_, thm') = Thm.varifyT' [] thm;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
    99
    val tvars = Term.add_tvars (Thm.prop_of thm') [];
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   100
    val unconstr = map (Thm.ctyp_of thy o TVar) tvars;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   101
    val instmap = map2 (fn (v_i, _) => fn (v, sort) => pairself (Thm.ctyp_of thy)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   102
      (TVar (v_i, []), TFree (v, sort))) tvars tfrees;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   103
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   104
    thm'
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   105
    |> fold Thm.unconstrainT unconstr
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   106
    |> Thm.instantiate (instmap, [])
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   107
    |> Tactic.rule_by_tactic ((REPEAT o CHANGED o ALLGOALS o Tactic.resolve_tac) (AxClass.class_intros thy))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   108
  end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   109
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   110
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   111
(** graph algorithm **)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   112
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   113
val timing = ref false;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   114
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   115
local
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   116
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   117
exception INVALID of CodeUnit.const list * string;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   118
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   119
fun resort_thms algebra tap_typ [] = []
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   120
  | resort_thms algebra tap_typ (thms as thm :: _) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   121
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   122
        val thy = Thm.theory_of_thm thm;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   123
        val cs = fold_consts (insert (op =)) thms [];
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   124
        fun match_const c (ty, ty_decl) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   125
          let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   126
            val tys = CodeUnit.typargs thy (c, ty);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   127
            val sorts = map (snd o dest_TVar) (CodeUnit.typargs thy (c, ty_decl));
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   128
          in fold2 (curry (CodeUnit.typ_sort_inst algebra)) tys sorts end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   129
        fun match (c_ty as (c, ty)) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   130
          case tap_typ c_ty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   131
           of SOME ty_decl => match_const c (ty, ty_decl)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   132
            | NONE => I;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   133
        val tvars = fold match cs Vartab.empty;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   134
      in map (CodeUnit.inst_thm tvars) thms end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   135
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   136
fun resort_funcss thy algebra funcgr =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   137
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   138
    val typ_funcgr = try (fst o Constgraph.get_node funcgr o CodeUnit.const_of_cexpr thy);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   139
    fun resort_dep (const, thms) = (const, resort_thms algebra typ_funcgr thms)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   140
      handle Sorts.CLASS_ERROR e => raise INVALID ([const], Sorts.msg_class_error (Sign.pp thy) e
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   141
                    ^ ",\nfor constant " ^ CodeUnit.string_of_const thy const
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   142
                    ^ "\nin defining equations\n"
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   143
                    ^ (cat_lines o map string_of_thm) thms)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   144
    fun resort_rec tap_typ (const, []) = (true, (const, []))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   145
      | resort_rec tap_typ (const, thms as thm :: _) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   146
          let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   147
            val (_, ty) = CodeUnit.head_func thm;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   148
            val thms' as thm' :: _ = resort_thms algebra tap_typ thms
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   149
            val (_, ty') = CodeUnit.head_func thm';
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   150
          in (Sign.typ_equiv thy (ty, ty'), (const, thms')) end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   151
    fun resort_recs funcss =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   152
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   153
        fun tap_typ c_ty = case try (CodeUnit.const_of_cexpr thy) c_ty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   154
         of SOME const => AList.lookup (CodeUnit.eq_const) funcss const
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   155
              |> these
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   156
              |> try hd
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   157
              |> Option.map (snd o CodeUnit.head_func)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   158
          | NONE => NONE;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   159
        val (unchangeds, funcss') = split_list (map (resort_rec tap_typ) funcss);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   160
        val unchanged = fold (fn x => fn y => x andalso y) unchangeds true;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   161
      in (unchanged, funcss') end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   162
    fun resort_rec_until funcss =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   163
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   164
        val (unchanged, funcss') = resort_recs funcss;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   165
      in if unchanged then funcss' else resort_rec_until funcss' end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   166
  in map resort_dep #> resort_rec_until end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   167
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   168
fun instances_of thy algebra insts =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   169
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   170
    val thy_classes = (#classes o Sorts.rep_algebra o Sign.classes_of) thy;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   171
    fun all_classops tyco class =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   172
      try (AxClass.params_of_class thy) class
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   173
      |> Option.map snd
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   174
      |> these
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   175
      |> map (fn (c, _) => (c, SOME tyco))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   176
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   177
    Symtab.empty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   178
    |> fold (fn (tyco, class) =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   179
        Symtab.map_default (tyco, []) (insert (op =) class)) insts
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   180
    |> (fn tab => Symtab.fold (fn (tyco, classes) => append (maps (all_classops tyco)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   181
         (Graph.all_succs thy_classes classes))) tab [])
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   182
  end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   183
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   184
fun instances_of_consts thy algebra funcgr consts =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   185
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   186
    fun inst (cexpr as (c, ty)) = insts_of thy algebra c
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   187
      ((fst o Constgraph.get_node funcgr o CodeUnit.const_of_cexpr thy) cexpr)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   188
      ty handle CLASS_ERROR => [];
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   189
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   190
    []
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   191
    |> fold (fold (insert (op =)) o inst) consts
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   192
    |> instances_of thy algebra
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   193
  end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   194
24283
haftmann
parents: 24219
diff changeset
   195
fun ensure_const' thy algebra funcgr const auxgr =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   196
  if can (Constgraph.get_node funcgr) const
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   197
    then (NONE, auxgr)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   198
  else if can (Constgraph.get_node auxgr) const
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   199
    then (SOME const, auxgr)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   200
  else if is_some (Code.get_datatype_of_constr thy const) then
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   201
    auxgr
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   202
    |> Constgraph.new_node (const, [])
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   203
    |> pair (SOME const)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   204
  else let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   205
    val thms = Code.these_funcs thy const
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   206
      |> CodeUnit.norm_args
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   207
      |> CodeUnit.norm_varnames CodeName.purify_tvar CodeName.purify_var;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   208
    val rhs = consts_of (const, thms);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   209
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   210
    auxgr
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   211
    |> Constgraph.new_node (const, thms)
24283
haftmann
parents: 24219
diff changeset
   212
    |> fold_map (ensure_const thy algebra funcgr) rhs
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   213
    |-> (fn rhs' => fold (fn SOME const' => Constgraph.add_edge (const, const')
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   214
                           | NONE => I) rhs')
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   215
    |> pair (SOME const)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   216
  end
24283
haftmann
parents: 24219
diff changeset
   217
and ensure_const thy algebra funcgr const =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   218
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   219
    val timeap = if !timing
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   220
      then Output.timeap_msg ("time for " ^ CodeUnit.string_of_const thy const)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   221
      else I;
24283
haftmann
parents: 24219
diff changeset
   222
  in timeap (ensure_const' thy algebra funcgr const) end;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   223
24283
haftmann
parents: 24219
diff changeset
   224
fun merge_funcss thy algebra raw_funcss funcgr =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   225
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   226
    val funcss = raw_funcss
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   227
      |> resort_funcss thy algebra funcgr
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   228
      |> filter_out (can (Constgraph.get_node funcgr) o fst);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   229
    fun typ_func const [] = Code.default_typ thy const
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   230
      | typ_func (_, NONE) (thm :: _) = (snd o CodeUnit.head_func) thm
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   231
      | typ_func (const as (c, SOME tyco)) (thms as (thm :: _)) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   232
          let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   233
            val (_, ty) = CodeUnit.head_func thm;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   234
            val SOME class = AxClass.class_of_param thy c;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   235
            val sorts_decl = Sorts.mg_domain algebra tyco [class];
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   236
            val tys = CodeUnit.typargs thy (c, ty);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   237
            val sorts = map (snd o dest_TVar) tys;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   238
          in if sorts = sorts_decl then ty
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   239
            else raise INVALID ([const], "Illegal instantation for class operation "
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   240
              ^ CodeUnit.string_of_const thy const
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   241
              ^ "\nin defining equations\n"
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   242
              ^ (cat_lines o map string_of_thm) thms)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   243
          end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   244
    fun add_funcs (const, thms) =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   245
      Constgraph.new_node (const, (typ_func const thms, thms));
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   246
    fun add_deps (funcs as (const, thms)) funcgr =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   247
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   248
        val deps = consts_of funcs;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   249
        val insts = instances_of_consts thy algebra funcgr
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   250
          (fold_consts (insert (op =)) thms []);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   251
      in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   252
        funcgr
24283
haftmann
parents: 24219
diff changeset
   253
        |> ensure_consts' thy algebra insts
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   254
        |> fold (curry Constgraph.add_edge const) deps
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   255
        |> fold (curry Constgraph.add_edge const) insts
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   256
       end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   257
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   258
    funcgr
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   259
    |> fold add_funcs funcss
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   260
    |> fold add_deps funcss
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   261
  end
24283
haftmann
parents: 24219
diff changeset
   262
and ensure_consts' thy algebra cs funcgr =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   263
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   264
    val auxgr = Constgraph.empty
24283
haftmann
parents: 24219
diff changeset
   265
      |> fold (snd oo ensure_const thy algebra funcgr) cs;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   266
  in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   267
    funcgr
24283
haftmann
parents: 24219
diff changeset
   268
    |> fold (merge_funcss thy algebra)
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   269
         (map (AList.make (Constgraph.get_node auxgr))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   270
         (rev (Constgraph.strong_conn auxgr)))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   271
  end handle INVALID (cs', msg)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   272
    => raise INVALID (fold (insert CodeUnit.eq_const) cs' cs, msg);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   273
24283
haftmann
parents: 24219
diff changeset
   274
fun ensure_consts thy consts funcgr =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   275
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   276
    val algebra = Code.coregular_algebra thy
24283
haftmann
parents: 24219
diff changeset
   277
  in ensure_consts' thy algebra consts funcgr
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   278
    handle INVALID (cs', msg) => error (msg ^ ",\nwhile preprocessing equations for constant(s) "
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   279
    ^ commas (map (CodeUnit.string_of_const thy) cs'))
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   280
  end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   281
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   282
in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   283
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   284
(** retrieval interfaces **)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   285
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   286
val ensure_consts = ensure_consts;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   287
24283
haftmann
parents: 24219
diff changeset
   288
fun check_consts thy consts funcgr =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   289
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   290
    val algebra = Code.coregular_algebra thy;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   291
    fun try_const const funcgr =
24283
haftmann
parents: 24219
diff changeset
   292
      (SOME const, ensure_consts' thy algebra [const] funcgr)
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   293
      handle INVALID (cs', msg) => (NONE, funcgr);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   294
    val (consts', funcgr') = fold_map try_const consts funcgr;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   295
  in (map_filter I consts', funcgr') end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   296
24283
haftmann
parents: 24219
diff changeset
   297
fun ensure_consts_term_proto thy f ct funcgr =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   298
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   299
    fun consts_of thy t =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   300
      fold_aterms (fn Const c => cons (CodeUnit.const_of_cexpr thy c) | _ => I) t []
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   301
    fun rhs_conv conv thm =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   302
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   303
        val thm' = (conv o Thm.rhs_of) thm;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   304
      in Thm.transitive thm thm' end
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   305
    val _ = Sign.no_vars (Sign.pp thy) (Thm.term_of ct);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   306
    val _ = Term.fold_types (Type.no_tvars #> K I) (Thm.term_of ct) ();
24283
haftmann
parents: 24219
diff changeset
   307
    val thm1 = Code.preprocess_conv ct;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   308
    val ct' = Thm.rhs_of thm1;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   309
    val consts = consts_of thy (Thm.term_of ct');
24283
haftmann
parents: 24219
diff changeset
   310
    val funcgr' = ensure_consts thy consts funcgr;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   311
    val algebra = Code.coregular_algebra thy;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   312
    val (_, thm2) = Thm.varifyT' [] thm1;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   313
    val thm3 = Thm.reflexive (Thm.rhs_of thm2);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   314
    val typ_funcgr = try (fst o Constgraph.get_node funcgr' o CodeUnit.const_of_cexpr thy);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   315
    val [thm4] = resort_thms algebra typ_funcgr [thm3];
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   316
    val tfrees = Term.add_tfrees (Thm.prop_of thm1) [];
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   317
    fun inst thm =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   318
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   319
        val tvars = Term.add_tvars (Thm.prop_of thm) [];
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   320
        val instmap = map2 (fn (v_i, sort) => fn (v, _) => pairself (Thm.ctyp_of thy)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   321
          (TVar (v_i, sort), TFree (v, sort))) tvars tfrees;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   322
      in Thm.instantiate (instmap, []) thm end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   323
    val thm5 = inst thm2;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   324
    val thm6 = inst thm4;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   325
    val ct'' = Thm.rhs_of thm6;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   326
    val cs = fold_aterms (fn Const c => cons c | _ => I) (Thm.term_of ct'') [];
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   327
    val drop = drop_classes thy tfrees;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   328
    val instdefs = instances_of_consts thy algebra funcgr' cs;
24283
haftmann
parents: 24219
diff changeset
   329
    val funcgr'' = ensure_consts thy instdefs funcgr';
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   330
  in (f funcgr'' drop ct'' thm5, funcgr'') end;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   331
24283
haftmann
parents: 24219
diff changeset
   332
fun ensure_consts_eval thy conv =
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   333
  let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   334
    fun conv' funcgr drop_classes ct thm1 =
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   335
      let
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   336
        val thm2 = conv funcgr ct;
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   337
        val thm3 = Code.postprocess_conv (Thm.rhs_of thm2);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   338
        val thm23 = drop_classes (Thm.transitive thm2 thm3);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   339
      in
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   340
        Thm.transitive thm1 thm23 handle THM _ =>
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   341
          error ("eval_conv - could not construct proof:\n"
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   342
          ^ (cat_lines o map string_of_thm) [thm1, thm2, thm3])
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   343
      end;
24283
haftmann
parents: 24219
diff changeset
   344
  in ensure_consts_term_proto thy conv' end;
haftmann
parents: 24219
diff changeset
   345
haftmann
parents: 24219
diff changeset
   346
fun ensure_consts_term thy f =
haftmann
parents: 24219
diff changeset
   347
  let
haftmann
parents: 24219
diff changeset
   348
    fun f' funcgr drop_classes ct thm1 = f funcgr ct;
haftmann
parents: 24219
diff changeset
   349
  in ensure_consts_term_proto thy f' end;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   350
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   351
end; (*local*)
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   352
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   353
structure Funcgr = CodeDataFun
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   354
(struct
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   355
  type T = T;
24283
haftmann
parents: 24219
diff changeset
   356
  val empty = Constgraph.empty;
haftmann
parents: 24219
diff changeset
   357
  fun merge _ _ = Constgraph.empty;
haftmann
parents: 24219
diff changeset
   358
  fun purge _ NONE _ = Constgraph.empty
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   359
    | purge _ (SOME cs) funcgr =
24283
haftmann
parents: 24219
diff changeset
   360
        Constgraph.del_nodes ((Constgraph.all_preds funcgr 
haftmann
parents: 24219
diff changeset
   361
          o filter (can (Constgraph.get_node funcgr))) cs) funcgr;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   362
end);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   363
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   364
fun make thy =
24283
haftmann
parents: 24219
diff changeset
   365
  Funcgr.change thy o ensure_consts thy;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   366
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   367
fun make_consts thy =
24283
haftmann
parents: 24219
diff changeset
   368
  Funcgr.change_yield thy o check_consts thy;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   369
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   370
fun eval_conv thy f =
24283
haftmann
parents: 24219
diff changeset
   371
  fst o Funcgr.change_yield thy o ensure_consts_eval thy f;
haftmann
parents: 24219
diff changeset
   372
haftmann
parents: 24219
diff changeset
   373
fun eval_term thy f =
haftmann
parents: 24219
diff changeset
   374
  fst o Funcgr.change_yield thy o ensure_consts_term thy f;
24219
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   375
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   376
fun intervene thy funcgr = Funcgr.change thy (K funcgr);
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   377
e558fe311376 new structure for code generator modules
haftmann
parents:
diff changeset
   378
end; (*struct*)