TFL/thry.sml
author paulson
Thu, 08 Jul 1999 13:38:41 +0200
changeset 6915 4ab8e31a8421
parent 6397 e70ae9b575cc
child 8416 8eb32cd3122e
permissions -rw-r--r--
Now if_weak_cong is a standard congruence rule
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3302
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3245
diff changeset
     1
(*  Title:      TFL/thry
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
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     7
structure Thry : Thry_sig (* LThry_sig *) = 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     8
struct
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     9
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    10
structure S = USyntax;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    11
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    12
fun THRY_ERR{func,mesg} = Utils.ERR{module = "Thry",func=func,mesg=mesg};
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    13
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    14
(*---------------------------------------------------------------------------
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    15
 *    Matching 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    16
 *---------------------------------------------------------------------------*)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    17
3353
9112a2efb9a3 Removal of module Mask and datatype binding with its constructor |->
paulson
parents: 3332
diff changeset
    18
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
    19
      fun tmbind (x,y) = (Var  (x,type_of y), y)
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    20
in
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    21
 fun match_term thry pat ob = 
6397
e70ae9b575cc Theory.sign_of;
wenzelm
parents: 5193
diff changeset
    22
    let val tsig = #tsig(Sign.rep_sg(Theory.sign_of thry))
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    23
        val (ty_theta,tm_theta) = Pattern.match tsig (pat,ob)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    24
    in (map tmbind tm_theta, map tybind ty_theta)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    25
    end
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    26
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    27
 fun match_type thry pat ob = 
6397
e70ae9b575cc Theory.sign_of;
wenzelm
parents: 5193
diff changeset
    28
    map tybind(Type.typ_match (#tsig(Sign.rep_sg(Theory.sign_of thry))) ([],(pat,ob)))
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    29
end;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    30
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    31
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    32
(*---------------------------------------------------------------------------
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    33
 * Typing 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    34
 *---------------------------------------------------------------------------*)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    35
6397
e70ae9b575cc Theory.sign_of;
wenzelm
parents: 5193
diff changeset
    36
fun typecheck thry = Thm.cterm_of (Theory.sign_of thry);
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    37
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    38
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    39
(*---------------------------------------------------------------------------
5193
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    40
 * Get information about datatypes
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    41
 *---------------------------------------------------------------------------*)
4107
2270829d2364 adapted to new datatypes thy info;
wenzelm
parents: 3405
diff changeset
    42
5193
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    43
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
    44
5193
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    45
fun match_info thy tname =
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    46
  case (DatatypePackage.case_const_of thy tname, DatatypePackage.constrs_of thy tname) of
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    47
      (Some case_const, Some constructors) =>
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    48
        Some {case_const = case_const, constructors = constructors}
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    49
    | _ => None;
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    50
5193
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    51
fun induct_info thy tname = case get_info thy tname of
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    52
        None => None
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    53
      | Some {nchotomy, ...} =>
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    54
          Some {nchotomy = nchotomy,
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    55
                constructors = the (DatatypePackage.constrs_of thy tname)};
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    56
5193
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    57
fun extract_info thy =
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    58
 let val infos = map snd (Symtab.dest (DatatypePackage.get_datatypes thy))
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    59
 in {case_congs = map (mk_meta_eq o #case_cong) infos,
5f6f7195dacf Adapted to new datatype package.
berghofe
parents: 4107
diff changeset
    60
     case_rewrites = flat (map (map mk_meta_eq o #case_rewrites) infos)}
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    61
 end;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    62
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    63
end; (* Thry *)