TFL/thry.sml
author nipkow
Thu, 26 Oct 2000 14:59:38 +0200
changeset 10343 24c87e1366d8
parent 9867 bf8300fa4238
permissions -rw-r--r--
*** empty log message ***
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9867
wenzelm
parents: 8416
diff changeset
     1
(*  Title:      TFL/thry.sml
3302
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3245
diff changeset
     2
    ID:         $Id$
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3245
diff changeset
     3
    Author:     Konrad Slind, Cambridge University Computer Laboratory
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3245
diff changeset
     4
    Copyright   1997  University of Cambridge
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3245
diff changeset
     5
*)
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3245
diff changeset
     6
9867
wenzelm
parents: 8416
diff changeset
     7
signature Thry_sig =
wenzelm
parents: 8416
diff changeset
     8
sig
wenzelm
parents: 8416
diff changeset
     9
  val match_term : theory -> term -> term 
wenzelm
parents: 8416
diff changeset
    10
                    -> (term*term)list * (typ*typ)list
wenzelm
parents: 8416
diff changeset
    11
wenzelm
parents: 8416
diff changeset
    12
  val match_type : theory -> typ -> typ -> (typ*typ)list
wenzelm
parents: 8416
diff changeset
    13
wenzelm
parents: 8416
diff changeset
    14
  val typecheck : theory -> term -> cterm
wenzelm
parents: 8416
diff changeset
    15
wenzelm
parents: 8416
diff changeset
    16
  (* Datatype facts of various flavours *)
wenzelm
parents: 8416
diff changeset
    17
  val match_info: theory -> string -> {constructors:term list,
wenzelm
parents: 8416
diff changeset
    18
                                     case_const:term} option
wenzelm
parents: 8416
diff changeset
    19
wenzelm
parents: 8416
diff changeset
    20
  val induct_info: theory -> string -> {constructors:term list,
wenzelm
parents: 8416
diff changeset
    21
                                      nchotomy:thm} option
wenzelm
parents: 8416
diff changeset
    22
wenzelm
parents: 8416
diff changeset
    23
  val extract_info: theory -> {case_congs:thm list, case_rewrites:thm list}
wenzelm
parents: 8416
diff changeset
    24
wenzelm
parents: 8416
diff changeset
    25
end;
wenzelm
parents: 8416
diff changeset
    26
wenzelm
parents: 8416
diff changeset
    27
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    28
structure Thry : Thry_sig (* LThry_sig *) = 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    29
struct
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    30
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    31
structure S = USyntax;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    32
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    33
fun THRY_ERR{func,mesg} = Utils.ERR{module = "Thry",func=func,mesg=mesg};
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    34
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    35
(*---------------------------------------------------------------------------
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    36
 *    Matching 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    37
 *---------------------------------------------------------------------------*)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    38
3353
9112a2efb9a3 Removal of module Mask and datatype binding with its constructor |->
paulson
parents: 3332
diff changeset
    39
local fun tybind (x,y) = (TVar (x,["term"]) , y)
9112a2efb9a3 Removal of module Mask and datatype binding with its constructor |->
paulson
parents: 3332
diff changeset
    40
      fun tmbind (x,y) = (Var  (x,type_of y), y)
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    41
in
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    42
 fun match_term thry pat ob = 
6397
e70ae9b575cc Theory.sign_of;
wenzelm
parents: 5193
diff changeset
    43
    let val tsig = #tsig(Sign.rep_sg(Theory.sign_of thry))
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    44
        val (ty_theta,tm_theta) = Pattern.match tsig (pat,ob)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    45
    in (map tmbind tm_theta, map tybind ty_theta)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    46
    end
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    47
8416
8eb32cd3122e Type.typ_match now uses Vartab instead of association lists.
berghofe
parents: 6397
diff changeset
    48
 fun match_type thry pat ob = map tybind (Vartab.dest
8eb32cd3122e Type.typ_match now uses Vartab instead of association lists.
berghofe
parents: 6397
diff changeset
    49
   (Type.typ_match (#tsig(Sign.rep_sg(Theory.sign_of thry))) (Vartab.empty, (pat,ob))))
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    50
end;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    51
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    52
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    53
(*---------------------------------------------------------------------------
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    54
 * Typing 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    55
 *---------------------------------------------------------------------------*)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    56
6397
e70ae9b575cc Theory.sign_of;
wenzelm
parents: 5193
diff changeset
    57
fun typecheck thry = Thm.cterm_of (Theory.sign_of thry);
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    58
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    59
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    60
(*---------------------------------------------------------------------------
5193
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    61
 * Get information about datatypes
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    62
 *---------------------------------------------------------------------------*)
4107
2270829d2364 adapted to new datatypes thy info;
wenzelm
parents: 3405
diff changeset
    63
5193
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    64
fun get_info thy ty = Symtab.lookup (DatatypePackage.get_datatypes thy, ty);
4107
2270829d2364 adapted to new datatypes thy info;
wenzelm
parents: 3405
diff changeset
    65
5193
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    66
fun match_info thy tname =
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    67
  case (DatatypePackage.case_const_of thy tname, DatatypePackage.constrs_of thy tname) of
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    68
      (Some case_const, Some constructors) =>
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    69
        Some {case_const = case_const, constructors = constructors}
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    70
    | _ => None;
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    71
5193
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    72
fun induct_info thy tname = case get_info thy tname of
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    73
        None => None
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    74
      | Some {nchotomy, ...} =>
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    75
          Some {nchotomy = nchotomy,
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    76
                constructors = the (DatatypePackage.constrs_of thy tname)};
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    77
5193
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    78
fun extract_info thy =
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    79
 let val infos = map snd (Symtab.dest (DatatypePackage.get_datatypes thy))
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    80
 in {case_congs = map (mk_meta_eq o #case_cong) infos,
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    81
     case_rewrites = flat (map (map mk_meta_eq o #case_rewrites) infos)}
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    82
 end;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    83
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    84
end; (* Thry *)