src/Tools/Metis/src/LiteralNet.sig
author blanchet
Mon, 13 Sep 2010 21:09:43 +0200
changeset 39348 6f9c9899f99f
child 39349 2d0a4361c3ef
permissions -rw-r--r--
new version of the Metis files
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     1
(* ========================================================================= *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     2
(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS           *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     3
(* Copyright (c) 2001-2006 Joe Hurd, distributed under the GNU GPL version 2 *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     4
(* ========================================================================= *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     5
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     6
signature LiteralNet =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     7
sig
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     8
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     9
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    10
(* A type of literal sets that can be efficiently matched and unified.       *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    11
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    12
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    13
type parameters = {fifo : bool}
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    14
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    15
type 'a literalNet
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    16
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    17
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    18
(* Basic operations.                                                         *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    19
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    20
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    21
val new : parameters -> 'a literalNet
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    22
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    23
val size : 'a literalNet -> int
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    24
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    25
val profile : 'a literalNet -> {positive : int, negative : int}
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    26
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    27
val insert : 'a literalNet -> Literal.literal * 'a -> 'a literalNet
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    28
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    29
val fromList : parameters -> (Literal.literal * 'a) list -> 'a literalNet
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    30
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    31
val filter : ('a -> bool) -> 'a literalNet -> 'a literalNet
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    32
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    33
val toString : 'a literalNet -> string
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    34
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    35
val pp : 'a Print.pp -> 'a literalNet Print.pp
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    36
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    37
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    38
(* Matching and unification queries.                                         *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    39
(*                                                                           *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    40
(* These function return OVER-APPROXIMATIONS!                                *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    41
(* Filter afterwards to get the precise set of satisfying values.            *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    42
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    43
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    44
val match : 'a literalNet -> Literal.literal -> 'a list
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    45
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    46
val matched : 'a literalNet -> Literal.literal -> 'a list
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    47
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    48
val unify : 'a literalNet -> Literal.literal -> 'a list
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    49
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    50
end