Nominal/nominal_dt_data.ML
author Christian Urban <urbanc@in.tum.de>
Sat, 02 Jul 2011 00:27:47 +0100
changeset 2931 aaef9dec5e1d
parent 2927 116f9ba4f59f
child 2957 01ff621599bc
permissions -rw-r--r--
side-by-side tests of lets with single assignment; deep-binder case works if the recursion is avoided using an auxiliary function

(* Author: Christian Urban

   data about nominal datatypes 
*)

signature NOMINAL_DT_DATA =
sig
  type info =
   {inject : thm list,
    distinct : thm list,
    strong_inducts : thm list,
    strong_exhaust : thm list}

  val get_all_info: Proof.context -> (string * info) list
  val get_info: Proof.context -> string -> info option
  val the_info: Proof.context -> string -> info 
  val register_info: (string * info) -> Context.generic -> Context.generic
  val mk_infos: string list -> thm list -> thm list -> thm list -> thm list -> (string * info) list

  datatype alpha_result = AlphaResult of
    {alpha_names      : string list,
     alpha_trms       : term list,
     alpha_tys        : typ list,
     alpha_bn_names   : string list,
     alpha_bn_trms    : term list,
     alpha_bn_tys     : typ list,
     alpha_intros     : thm list,
     alpha_cases      : thm list,
     alpha_raw_induct : thm}
   
end

structure Nominal_Dt_Data: NOMINAL_DT_DATA =
struct


(* information generated by nominal_datatype *)
type info =
   {inject : thm list,
    distinct : thm list,
    strong_inducts : thm list,
    strong_exhaust : thm list}

structure NominalData = Generic_Data
 (type T = info Symtab.table
  val empty = Symtab.empty
  val extend = I
  val merge = Symtab.merge (K true))

val get_all_info = Symtab.dest o NominalData.get o Context.Proof
val get_info = Symtab.lookup o NominalData.get o Context.Proof
val register_info = NominalData.map o Symtab.update

fun the_info thy name =
  (case get_info thy name of
    SOME info => info
  | NONE => error ("Unknown nominal datatype " ^ quote name))

fun mk_infos ty_names inject distinct strong_inducts strong_exhaust =
let
  fun aux ty_name =
    (ty_name, {inject = inject,
               distinct = distinct,
               strong_inducts = strong_inducts,
               strong_exhaust = strong_exhaust 
              })
in
  map aux ty_names
end

datatype alpha_result = AlphaResult of
  {alpha_names      : string list,
   alpha_trms       : term list,
   alpha_tys        : typ list,
   alpha_bn_names   : string list,
   alpha_bn_trms    : term list,
   alpha_bn_tys     : typ list,
   alpha_intros     : thm list,
   alpha_cases      : thm list,
   alpha_raw_induct : thm}


end