src/Tools/Argo/argo_lit.ML
author nipkow
Sun, 06 May 2018 13:51:37 +0200
changeset 68083 d730a8cfc6e0
parent 63960 3daf02070be5
permissions -rw-r--r--
removed asm "finite"
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     1
(*  Title:      Tools/Argo/argo_lit.ML
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     2
    Author:     Sascha Boehme
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     3
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     4
Representation of literals. Literals are terms with a polarity, either positive or negative.
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     5
A literal for term t with positive polarity is equivalent to t.
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     6
A literal for term t with negative polarity is equivalent to the propositional negation of t.
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     7
*)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     8
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     9
signature ARGO_LIT =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    10
sig
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    11
  datatype literal = Pos of Argo_Term.term | Neg of Argo_Term.term
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    12
  val literal: Argo_Term.term -> bool -> literal
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    13
  val dest: literal -> Argo_Term.term * bool
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    14
  val term_of: literal -> Argo_Term.term
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    15
  val signed_id_of: literal -> int
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    16
  val signed_expr_of: literal -> Argo_Expr.expr
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    17
  val negate: literal -> literal
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    18
  val eq_id: literal * literal -> bool
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    19
  val eq_lit: literal * literal -> bool
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    20
  val dual_lit: literal * literal -> bool
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    21
end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    22
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    23
structure Argo_Lit: ARGO_LIT =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    24
struct
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    25
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    26
(* data type *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    27
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    28
datatype literal = Pos of Argo_Term.term | Neg of Argo_Term.term
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    29
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    30
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    31
(* operations *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    32
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    33
fun literal t true = Pos t
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    34
  | literal t false = Neg t
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    35
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    36
fun dest (Pos t) = (t, true)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    37
  | dest (Neg t) = (t, false)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    38
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    39
fun term_of (Pos t) = t
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    40
  | term_of (Neg t) = t
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    41
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    42
fun signed_id_of (Pos t) = Argo_Term.id_of t
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    43
  | signed_id_of (Neg t) = ~(Argo_Term.id_of t)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    44
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    45
fun signed_expr_of (Pos t) = Argo_Term.expr_of t
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    46
  | signed_expr_of (Neg t) = Argo_Expr.mk_not (Argo_Term.expr_of t)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    47
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    48
fun id_of (Pos t) = Argo_Term.id_of t
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    49
  | id_of (Neg t) = Argo_Term.id_of t
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    50
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    51
fun negate (Pos t) = Neg t
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    52
  | negate (Neg t) = Pos t
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    53
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    54
fun eq_id (lit1, lit2) = (id_of lit1 = id_of lit2)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    55
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    56
fun eq_lit (Pos t1, Pos t2) = Argo_Term.eq_term (t1, t2)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    57
  | eq_lit (Neg t1, Neg t2) = Argo_Term.eq_term (t1, t2)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    58
  | eq_lit _ = false
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    59
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    60
fun dual_lit (Pos t1, Neg t2) = Argo_Term.eq_term (t1, t2)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    61
  | dual_lit (Neg t1, Pos t2) = Argo_Term.eq_term (t1, t2)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    62
  | dual_lit _ = false
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    63
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    64
end