TFL/thry.sig
author paulson
Thu, 22 May 1997 15:13:16 +0200
changeset 3302 404fe31fd8d2
parent 3245 241838c01caf
child 3353 9112a2efb9a3
permissions -rw-r--r--
New headers and other minor changes
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
signature Thry_sig =
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     8
sig
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     9
  type 'a binding
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    10
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    11
  structure USyntax : USyntax_sig
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 2112
diff changeset
    12
  val match_term : theory -> term -> term 
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 2112
diff changeset
    13
                    -> term binding list * typ binding list
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    14
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 2112
diff changeset
    15
  val match_type : theory -> typ -> typ -> typ binding list
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    16
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 2112
diff changeset
    17
  val typecheck : theory -> term -> cterm
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    18
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 2112
diff changeset
    19
  val make_definition: theory -> string -> term -> thm * theory
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    20
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    21
  (* Datatype facts of various flavours *)
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 2112
diff changeset
    22
  val match_info: theory -> string -> {constructors:term list,
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 2112
diff changeset
    23
                                     case_const:term} option
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    24
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 2112
diff changeset
    25
  val induct_info: theory -> string -> {constructors:term list,
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 2112
diff changeset
    26
                                      nchotomy:thm} option
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    27
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 2112
diff changeset
    28
  val extract_info: theory -> {case_congs:thm list, case_rewrites:thm list}
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    29
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    30
end;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    31
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    32