TFL/tfl.sig
author paulson
Tue, 20 May 1997 11:49:57 +0200
changeset 3245 241838c01caf
parent 3191 14bd6e5985f1
child 3302 404fe31fd8d2
permissions -rw-r--r--
Removal of redundant code (unused or already present in Isabelle. This eliminates HOL compatibility but makes the code smaller and more readable
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     1
signature TFL_sig =
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     2
sig
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     3
   structure Rules: Rules_sig
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     4
   structure Thms : Thms_sig
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     5
   structure Thry : Thry_sig
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     6
   structure USyntax : USyntax_sig
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     7
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
     8
   datatype pattern = GIVEN of term * int
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
     9
                    | OMITTED of term * int
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    10
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    11
   val mk_functional : theory -> term
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    12
                       -> {functional:term,
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    13
                           pats: pattern list}
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    14
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    15
   val wfrec_definition0 : theory -> term -> term -> thm * theory
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
    16
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    17
   val post_definition : theory * (thm * pattern list)
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    18
                              -> {theory : theory,
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    19
                                  rules  : thm, 
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    20
                                    TCs  : term list list,
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    21
                          full_pats_TCs  : (term * term list) list, 
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
    22
                               patterns  : pattern list}
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
    23
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: 3191
diff changeset
    25
   val wfrec_eqns : theory -> term
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    26
                     -> {WFR : term, 
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    27
                         proto_def : term,
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    28
                         extracta :(thm * term list) list,
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    29
                         pats  : pattern list}
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    30
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    31
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    32
   val lazyR_def : theory
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    33
                   -> term
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    34
                   -> {theory:theory,
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    35
                       rules :thm,
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    36
                           R :term,
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    37
                       full_pats_TCs :(term * term list) list, 
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    38
                       patterns: pattern list}
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    39
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    40
   val mk_induction : theory 
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    41
                       -> term -> term 
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    42
                         -> (term * term list) list
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    43
                           -> thm
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    44
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    45
   val postprocess: {WFtac:tactic, terminator:tactic, simplifier:cterm -> thm}
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    46
                     -> theory 
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    47
                      -> {rules:thm, induction:thm, TCs:term list list}
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    48
                       -> {rules:thm, induction:thm, nested_tcs:thm list}
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    49
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    50
   structure Context
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    51
     : sig
3245
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    52
         val read : unit -> thm list
241838c01caf Removal of redundant code (unused or already present in Isabelle.
paulson
parents: 3191
diff changeset
    53
         val write : thm list -> unit
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    54
       end
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    55
end;