src/Tools/Metis/src/selftest.sml
author wenzelm
Wed, 17 Mar 2021 22:24:57 +0100
changeset 73446 d1c4c2395650
parent 72004 913162a47d9f
child 74358 6ab3116a251a
permissions -rw-r--r--
more Isabelle symbol definitions for Z Notation, based on https://github.com/isabelle-utp/Z_Toolkit 998c9f7880d3 by Simon Foster; NB: no bold version of 0x2900 due to fontforge crash "Internal Error: Some fragments did not join";
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
(* METIS TESTS                                                               *)
72004
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
     3
(* Copyright (c) 2004 Joe Leslie-Hurd, distributed under the BSD License     *)
39348
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
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
     7
(* Dummy versions of Moscow ML declarations to stop real compilers barfing.  *)
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
(*mlton
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    11
val quotation = ref true;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    12
val quietdec  = ref true;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    13
val loadPath  = ref ([] : string list);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    14
val load = fn (_ : string) => ();
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    15
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    16
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    17
(*polyml
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    18
val quotation = ref true;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    19
val quietdec  = ref true;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    20
val loadPath  = ref ([] : string list);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    21
val load = fn (_ : string) => ();
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    22
*)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    23
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    24
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    25
(* Load and open some useful modules                                         *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    26
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    27
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    28
val () = loadPath := !loadPath @ ["../bin/mosml"];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    29
val () = app load ["Options"];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    30
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    31
open Useful;
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 time = Portable.time;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    34
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    35
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    36
(* Problem data.                                                             *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    37
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    38
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    39
val ref oldquietdec = quietdec;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    40
val () = quietdec := true;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    41
val () = quotation := true;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    42
use "../src/problems.sml";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    43
val () = quietdec := oldquietdec;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    44
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    45
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    46
(* Helper functions.                                                         *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    47
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    48
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    49
fun partialOrderToString (SOME LESS) = "SOME LESS"
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    50
  | partialOrderToString (SOME GREATER) = "SOME GREATER"
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    51
  | partialOrderToString (SOME EQUAL) = "SOME EQUAL"
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    52
  | partialOrderToString NONE = "NONE";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    53
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    54
fun SAY s =
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    55
    TextIO.print
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    56
      ("-------------------------------------" ^
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    57
       "-------------------------------------\n" ^ s ^ "\n\n");
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    58
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    59
fun printval p x = (TextIO.print (Print.toString p x ^ "\n\n"); x);
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    60
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    61
fun mkCl p th = Clause.mk {parameters = p, id = Clause.newId (), thm = th};
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    62
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    63
val pvBool = printval Print.ppBool
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    64
and pvPo = printval (Print.ppMap partialOrderToString Print.ppString)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    65
and pvFm = printval Formula.pp
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    66
and pvFms = printval (Print.ppList Formula.pp)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    67
and pvThm = printval Thm.pp
72004
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
    68
and pvThms = printval (Print.ppList Thm.pp)
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    69
and pvEqn : Rule.equation -> Rule.equation = printval (Print.ppMap snd Thm.pp)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    70
and pvNet = printval (LiteralNet.pp Print.ppInt)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    71
and pvRw = printval Rewrite.pp
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    72
and pvU = printval Units.pp
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    73
and pvLits = printval LiteralSet.pp
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    74
and pvCl = printval Clause.pp
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    75
and pvCls = printval (Print.ppList Clause.pp)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    76
and pvM = printval Model.pp;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    77
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    78
val NV = Name.fromString
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    79
and NF = Name.fromString
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    80
and NR = Name.fromString;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    81
val V = Term.Var o NV
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    82
and C = (fn c => Term.Fn (NF c, []))
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    83
and T = Term.parse
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    84
and A = Atom.parse
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    85
and L = Literal.parse
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    86
and F = Formula.parse
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    87
and S = Subst.fromList;
42102
fcfd07f122d4 new version of Metis 2.3 (29 Dec. 2010)
blanchet
parents: 39502
diff changeset
    88
val LS = LiteralSet.fromList o List.map L;
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    89
val AX = Thm.axiom o LS;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    90
val CL = mkCl Clause.default o AX;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    91
val Q = (fn th => (Thm.destUnitEq th, th)) o AX o singleton
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    92
and U = (fn th => (Thm.destUnit th, th)) o AX o singleton;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    93
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    94
fun test_fun eq p r a =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
    95
  if eq r a then p a ^ "\n" else
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    96
    (TextIO.print
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    97
       ("\n\n" ^
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    98
        "test: should have\n-->" ^ p r ^ "<--\n\n" ^
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
    99
        "test: actually have\n-->" ^ p a ^ "<--\n\n");
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   100
     raise Fail "test: failed a test");
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   101
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
   102
fun test eq p r a = TextIO.print (test_fun eq p r a ^ "\n");
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   103
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   104
val test_tm = test Term.equal Term.toString o Term.parse;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   105
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   106
val test_fm = test Formula.equal Formula.toString o Formula.parse;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   107
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   108
fun test_id p f a = test p a (f a);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   109
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   110
fun chop_newline s =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   111
    if String.sub (s,0) = #"\n" then String.extract (s,1,NONE) else s;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   112
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   113
fun unquote (QUOTE q) = q
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   114
  | unquote (ANTIQUOTE _) = raise Fail "unquote";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   115
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   116
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   117
val () = SAY "The parser and pretty-printer";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   118
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   119
42102
fcfd07f122d4 new version of Metis 2.3 (29 Dec. 2010)
blanchet
parents: 39502
diff changeset
   120
fun prep l = (chop_newline o String.concat o List.map unquote) l;
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   121
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   122
fun mini_print n fm = withRef (Print.lineLength,n) Formula.toString fm;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   123
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   124
fun testlen_pp n q =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   125
    (fn s => test_fun equal I s ((mini_print n o Formula.fromString) s))
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   126
      (prep q);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   127
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
   128
fun test_pp q = TextIO.print (testlen_pp 40 q ^ "\n");
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   129
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   130
val () = test_pp `3 = f x`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   131
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   132
val () = test_pp `f x y = y`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   133
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   134
val () = test_pp `P x y`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   135
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   136
val () = test_pp `P (f x) y`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   137
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   138
val () = test_pp `f x = 3`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   139
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   140
val () = test_pp `!x. P x y`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   141
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   142
val () = test_pp `!x y. P x y`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   143
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   144
val () = test_pp `!x y z. P x y z`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   145
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   146
val () = test_pp `x = y`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   147
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   148
val () = test_pp `x = 3`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   149
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   150
val () = test_pp `x + y = y`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   151
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   152
val () = test_pp `x / y * z = w`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   153
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   154
val () = test_pp `x * y * z = x * (y * z)`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   155
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   156
val () = test_pp `!x. ?y. x <= y /\ y <= x`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   157
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   158
val () = test_pp `?x. !y. x + y = y /\ y <= x`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   159
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   160
val () = test_pp `p /\ q \/ r /\ p ==> q <=> p`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   161
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   162
val () = test_pp `p`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   163
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   164
val () = test_pp `~!x. bool x`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   165
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   166
val () = test_pp `p ==> !x. bool x`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   167
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   168
val () = test_pp `p ==> ~!x. bool x`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   169
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   170
val () = test_pp `~!x. bool x`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   171
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   172
val () = test_pp `~~!x. bool x`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   173
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   174
val () = test_pp `hello + there <> everybody`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   175
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   176
val () = test_pp `!x y. ?z. x < z /\ y < z`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   177
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   178
val () = test_pp `~(!x. P x) <=> ?y. ~P y`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   179
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   180
val () = test_pp `?y. x < y ==> !v. ?w. x * v < y * w`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   181
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   182
val () = test_pp `(<=)`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   183
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   184
val () = test_pp `(<=) <= b`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   185
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   186
val () = test_pp `(<=) <= (+)`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   187
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   188
val () = test_pp `(<=) x`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   189
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   190
val () = test_pp `(<=) <= (+) x`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   191
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   192
val () = test_pp `~B (P % ((,) % c_a % v_b))`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   193
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   194
val () = test_pp `B ((<=) % 0 % (LENGTH % NIL))`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   195
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   196
val () = test_pp `~(a = b)`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   197
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   198
val () = test_pp `!x. p x ==> !y. p y`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   199
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   200
val () = test_pp `(!x. p x) ==> !y. p y`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   201
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   202
val () = test_pp `!x. ~~x = x`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   203
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   204
val () = test_pp `x + (y + z) = a`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   205
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   206
val () = test_pp `(x @ y) @ z = a`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   207
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   208
val () = test_pp `p ((a @ a) @ a = a)`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   209
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   210
val () = test_pp `!x y z. (x @ y) @ z = (x @ z) @ y @ z`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   211
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   212
val () = test_pp `~(!x. q x) /\ p`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   213
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   214
val () = test_pp `!x. f (~~x) b (~c)`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   215
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   216
val () = test_pp `p ==> ~(a /\ b)`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   217
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   218
val () = test_pp `!water. drinks (water)`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   219
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   220
val () = test_pp `!vat water. drinks ((vat) p x (water))`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   221
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   222
val () = test_pp `!x y. ~{x < y} /\ T`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   223
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   224
val () = test_pp `[3]`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   225
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   226
val () = test_pp `
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   227
!x y z. ?x' y' z'.
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   228
  P x y z ==> P x' y' z'`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   229
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   230
val () = test_pp `
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   231
(!x. P x ==> !x. Q x) /\
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   232
((!x. Q x \/ R x) ==> ?x. Q x /\ R x) /\
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   233
((?x. R x) ==> !x. L x ==> M x) ==>
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   234
!x. P x /\ L x ==> M x`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   235
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   236
val () = test_pp `
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   237
!x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   238
  x12 x13 x14 x15 x16 x17 x18 x19 x20
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   239
  x21 x22 x23 x24 x25 x26 x27 x28 x29
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   240
  x30 x31 x32. ?y0 y1 y2 y3 y4 y5 y6 y7.
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   241
  P`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   242
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   243
val () = test_pp `
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   244
!x x x x x x x x x x x x x x x x x x x x
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   245
  x x x x x x x x x x. ?y y y y y y y y
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   246
  y y y y y y y y y y y.
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   247
  P (x, y) /\ P (x, y) /\ P (x, y) /\
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   248
  P (x, y) /\ P (x, y) /\ P (x, y) /\
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   249
  P (x, y) /\ P (x, y) /\ P (x, y) /\
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   250
  P (x, y) /\ P (x, y) /\ P (x, y) /\
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   251
  P (x, y) /\ P (x, y) /\
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   252
  ~~~~~~~~~~~~~f
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   253
                 (f (f (f x y) (f x y))
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   254
                    (f (f x y) (f x y)))
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   255
                 (f (f (f x y) (f x y))
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   256
                    (f (f x y) (f x y)))`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   257
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   258
val () = test_pp `
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   259
(!x.
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   260
   extremely__long__predicate__name) /\
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   261
F`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   262
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   263
val () = test_pp `
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   264
(!x. x = x) /\
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   265
(!x y. ~(x = y) \/ y = x) /\
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   266
(!x y z.
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   267
   ~(x = y) \/ ~(y = z) \/ x = z) /\
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   268
(!x y z. b . x . y . z = x . (y . z)) /\
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   269
(!x y. t . x . y = y . x) /\
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   270
(!x y z. ~(x = y) \/ x . z = y . z) /\
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   271
(!x y z. ~(x = y) \/ z . x = z . y) ==>
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   272
~(b . (b . (t . b) . b) . t . x . y .
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   273
  z = y . (x . z)) ==> F`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   274
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   275
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   276
val () = SAY "Substitution";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   277
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   278
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   279
val () =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   280
    test Name.equal Name.toString (NV"x")
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   281
      (Term.variantPrime (NameSet.fromList [NV"y",NV"z" ]) (NV"x"));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   282
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   283
val () =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   284
    test Name.equal Name.toString (NV"x'")
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   285
      (Term.variantPrime (NameSet.fromList [NV"x",NV"y" ]) (NV"x"));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   286
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   287
val () =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   288
    test Name.equal Name.toString (NV"x''")
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   289
      (Term.variantPrime (NameSet.fromList [NV"x",NV"x'"]) (NV"x"));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   290
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   291
val () =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   292
    test Name.equal Name.toString (NV"x")
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   293
      (Term.variantNum (NameSet.fromList [NV"y",NV"z"]) (NV"x"));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   294
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   295
val () =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   296
    test Name.equal Name.toString (NV"x0")
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   297
      (Term.variantNum (NameSet.fromList [NV"x",NV"y"]) (NV"x"));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   298
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   299
val () =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   300
    test Name.equal Name.toString (NV"x1")
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   301
      (Term.variantNum (NameSet.fromList [NV"x",NV"x0"]) (NV"x"));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   302
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   303
val () =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   304
    test_fm
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   305
      `!x. x = $z`
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   306
      (Formula.subst (S [(NV"y", V"z")]) (F`!x. x = $y`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   307
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   308
val () =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   309
    test_fm
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   310
      `!x'. x' = $x`
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   311
      (Formula.subst (S [(NV"y", V"x")]) (F`!x. x = $y`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   312
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   313
val () =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   314
    test_fm
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   315
      `!x' x''. x' = $x ==> x' = x''`
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   316
      (Formula.subst (S [(NV"y", V"x")])
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   317
         (F`!x x'. x = $y ==> x = x'`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   318
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   319
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   320
val () = SAY "Unification";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   321
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   322
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   323
fun unify_and_apply tm1 tm2 =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   324
    Subst.subst (Subst.unify Subst.empty tm1 tm2) tm1;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   325
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   326
val () = test_tm `c` (unify_and_apply (V"x") (C"c"));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   327
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   328
val () = test_tm `c` (unify_and_apply (C"c") (V"x"));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   329
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   330
val () =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   331
    test_tm
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   332
      `f c`
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   333
      (unify_and_apply
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   334
         (Term.Fn (NF"f", [V"x"]))
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   335
         (Term.Fn (NF"f", [C"c"])));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   336
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   337
val () = test_tm `f 0 0 0` (unify_and_apply (T`f 0 $x $x`) (T`f $y $y $z`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   338
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   339
fun f x y = (printval Subst.pp (Atom.unify Subst.empty x y); ());
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   340
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   341
val () = f (NR"P", [V"x"]) (NR"P", [V"x"]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   342
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   343
val () = f (NR"P", [V"x"]) (NR"P", [C"c"]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   344
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   345
val () = f (A`P c_x`) (A`P $x`);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   346
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   347
val () = f (A`q $x (f $x)`) (A`q $y $z`);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   349
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   350
val () = SAY "The logical kernel";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   351
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   352
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   353
val th0 = AX [`p`,`q`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   354
val th1 = AX [`~p`,`r`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   355
val th2 = Thm.resolve (L`p`) th0 th1;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   356
val _ = printval Proof.pp (Proof.proof th2);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   357
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   358
val th0 = Rule.relationCongruence Atom.eqRelation;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   359
val th1 =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   360
    Thm.subst (S [(NV"y0",T`$x`),(NV"y1",T`$y`),(NV"x1",T`$z`),(NV"x0",T`$x`)]) th0;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   361
val th2 = Thm.resolve (L`$x = $x`) Rule.reflexivity th1;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   362
val th3 = Rule.symNeq (L`~($z = $y)`) th2;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   363
val _ = printval Proof.pp (Proof.proof th3);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   364
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   365
(* Testing the elimination of redundancies in proofs *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   366
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   367
val th0 = Rule.reflexivity;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   368
val th1 = Thm.subst (S [(NV"x", Term.Fn (NF"f", [V"y"]))]) th0;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   369
val th2 = Thm.subst (S [(NV"y", C"c")]) th1;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   370
val _ = printval Proof.pp (Proof.proof th2);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   371
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   372
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   373
val () = SAY "Derived rules of inference";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   374
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   375
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   376
val th0 = pvThm (AX [`$x = a`, `f a = $x`, `~(a = b)`, `a = $x`,
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   377
                     `$x = a`, `a = $x`, `~(b = a)`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   378
val th1 = pvThm (Rule.removeSym th0);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   379
val th2 = pvThm (Rule.symEq (L`a = $x`) th0);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   380
val th3 = pvThm (Rule.symEq (L`f a = $x`) th0);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   381
val th5 = pvThm (Rule.symNeq (L`~(a = b)`) th0);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   382
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   383
(* Testing the rewrConv conversion *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   384
val (x_y as (x,y), eqTh) = pvEqn (Q`e * (i $z * $z) = e`);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   385
val tm = Term.Fn (NF"f",[x]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   386
val path : int list = [0];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   387
val reflTh = Thm.refl tm;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   388
val reflLit = Thm.destUnit reflTh;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   389
val th = Thm.equality reflLit (1 :: path) y;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   390
val th = Thm.resolve reflLit reflTh th;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   391
val th = pvThm (try (Thm.resolve (Literal.mkEq x_y) eqTh) th);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   392
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   393
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   394
val () = SAY "Discrimination nets for literals";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   395
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   396
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   397
val n = pvNet (LiteralNet.new {fifo = true});
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   398
val n = pvNet (LiteralNet.insert n (L`P (f c $x a)`, 1));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   399
val n = pvNet (LiteralNet.insert n (L`P (f c $y a)`, 2));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   400
val n = pvNet (LiteralNet.insert n (L`P (f c a a)`, 3));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   401
val n = pvNet (LiteralNet.insert n (L`P (f c b a)`, 4));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   402
val n = pvNet (LiteralNet.insert n (L`~Q`, 5));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   403
val n = pvNet (LiteralNet.insert n (L`~Q`, 6));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   404
val n = pvNet (LiteralNet.insert n (L`~Q`, 7));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   405
val n = pvNet (LiteralNet.insert n (L`~Q`, 8));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   406
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   407
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   408
val () = SAY "The Knuth-Bendix ordering on terms";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   409
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   410
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   411
val kboOrder = KnuthBendixOrder.default;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   412
val kboCmp = KnuthBendixOrder.compare kboOrder;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   413
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   414
val x = pvPo (kboCmp (T`f a`, T`g b`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   415
val x = pvPo (kboCmp (T`f a b`, T`g b`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   416
val x = pvPo (kboCmp (T`f $x`, T`g a`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   417
val x = pvPo (kboCmp (T`f a $x`, T`g $x`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   418
val x = pvPo (kboCmp (T`f $x`, T`g $x`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   419
val x = pvPo (kboCmp (T`f $x`, T`f $x`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   420
val x = pvPo (kboCmp (T`$x + $y`, T`$x + $x`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   421
val x = pvPo (kboCmp (T`$x + $y + $x`, T`$y + $x + $x`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   422
val x = pvPo (kboCmp (T`$x + $y + $x`, T`$y * $x + $x`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   423
val x = pvPo (kboCmp (T`a`, T`$x`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   424
val x = pvPo (kboCmp (T`f a`, T`$x`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   425
val x = pvPo (kboCmp (T`f $x (f $y $z)`, T`f (f $x $y) $z`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   426
val x = pvPo (kboCmp (T`f (g $x a)`, T`f (h a $x)`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   427
val x = pvPo (kboCmp (T`f (g a)`, T`f (h $x)`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   428
val x = pvPo (kboCmp (T`f (h a)`, T`f (g $x)`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   429
val x = pvPo (kboCmp (T`f $y`, T`f (g a b c)`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   430
val x = pvPo (kboCmp (T`$x * $y + $x * $z`, T`$x * ($y + $z)`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   431
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   432
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   433
val () = SAY "Rewriting";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   434
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   435
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   436
val eqnsToRw = Rewrite.addList (Rewrite.new kboCmp) o enumerate;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   437
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   438
val eqns = [Q`e * $x = $x`, Q`$x * e = $x`, Q`i $x * $x = e`, Q`$x * i $x = e`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   439
val ax = pvThm (AX [`e * (i $z * $z) = i e * e`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   440
val th = pvThm (Rewrite.orderedRewrite kboCmp eqns ax);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   441
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   442
val rw = pvRw (eqnsToRw eqns);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   443
val th = pvThm (snd (try (Rewrite.rewriteConv rw kboCmp) (T`e * e`)));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   444
val th = pvThm (snd (try (Rewrite.rewriteConv rw kboCmp) (T`e * (i $z * $z)`)));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   445
val th = pvThm (try (Rewrite.rewriteRule rw kboCmp) ax);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   446
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   447
(* Bug check: in this one a literal goes missing, due to the Resolve in Subst *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   448
val eqns = [Q`f a = a`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   449
val ax = pvThm (AX [`~(g (f a) = b)`, `~(f a = a)`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   450
val th = pvThm (try (Rewrite.orderedRewrite kboCmp eqns) ax);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   451
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   452
(* Bug check: term paths were not being reversed before use *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   453
val eqns = [Q`a = f a`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   454
val ax = pvThm (AX [`a <= f (f a)`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   455
val th = pvThm (try (Rewrite.orderedRewrite kboCmp eqns) ax);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   456
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   457
(* Bug check: Equality used to complain if the literal didn't exist *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   458
val eqns = [Q`b = f a`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   459
val ax = pvThm (AX [`~(f a = f a)`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   460
val th = pvThm (try (Rewrite.orderedRewrite kboCmp eqns) ax);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   461
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   462
(* Testing the rewriting with disequalities in the same clause *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   463
val ax = pvThm (AX [`~(a = b)`, `P a`, `P b`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   464
val th = pvThm (try (Rewrite.orderedRewrite kboCmp []) ax);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   465
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   466
val ax = pvThm (AX [`~(f $x = $x)`, `P (f a)`, `P (f $x)`, `P (f $y)`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   467
val th = pvThm (try (Rewrite.orderedRewrite kboCmp []) ax);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   468
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   469
val ax = pvThm
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   470
  (AX [`~(f (f (f (f (f $x)))) = $x)`,
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   471
          `~(f (f (f (f (f (f (f (f $x))))))) = $x)`,
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   472
          `P (f $x)`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   473
val th = pvThm (try (Rewrite.orderedRewrite kboCmp []) ax);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   474
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   475
(* Symmetry should yield a tautology on ground clauses *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   476
val ax = pvThm (AX [`~(a = b)`, `b = a`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   477
val th = pvThm (try (Rewrite.orderedRewrite kboCmp []) ax);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   478
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   479
(* Transitivity should yield a tautology on ground clauses *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   480
val ax = pvThm (AX [`~(a = b)`, `~(b = c)`, `a = c`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   481
val th = pvThm (try (Rewrite.orderedRewrite kboCmp []) ax);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   482
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   483
(* Extended transitivity should yield a tautology on ground clauses *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   484
val ax = pvThm (AX [`~(a = b)`, `~(b = c)`, `~(c = d)`, `a = d`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   485
val th = pvThm (try (Rewrite.orderedRewrite kboCmp []) ax);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   486
72004
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   487
(* Bug discovered by Michael Farber *)
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   488
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   489
val eqns = [Q`f $x = c`];
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   490
val ax = pvThm (AX [`~(f $y = g a b)`,`p (g a b)`]);
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   491
val th = pvThm (try (Rewrite.orderedRewrite kboCmp eqns) ax);
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   492
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   493
val eqns = [Q`even (numeral c) = d`,
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   494
            Q`f (numeral c) $x = $x`];
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   495
val ax = pvThm
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   496
  (AX [`~(even (numeral c) = even $y)`,
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   497
       `p (even (f (numeral c) $y))`]);
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   498
val th = pvThm (try (Rewrite.orderedRewrite kboCmp eqns) ax);
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   499
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   500
val eqns = [Q`even (numeral c) = d`,
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   501
            Q`f (numeral c) $x = $x`,
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   502
            Q`g a b = numeral c`];
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   503
val ax = pvThm
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   504
  (AX [`~(even (numeral c) = even $y)`,
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   505
       `p (even (f (g a b) $y))`]);
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   506
val th = pvThm (try (Rewrite.orderedRewrite kboCmp eqns) ax);
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   507
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   508
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   509
val () = SAY "Unit cache";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   510
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   511
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   512
val u = pvU (Units.add Units.empty (U`~p $x`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   513
val u = pvU (Units.add u (U`a = b`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   514
val _ = pvThm (Units.reduce u (AX [`p 0`,`~(b = a)`]));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   515
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   516
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   517
val () = SAY "Negation normal form";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   518
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   519
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   520
val nnf = Normalize.nnf;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   521
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   522
val _ = pvFm (nnf (F`p /\ ~p`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   523
val _ = pvFm (nnf (F`(!x. P x) ==> ((?y. Q y) <=> (?z. P z /\ Q z))`));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   524
val _ = pvFm (nnf (F`~(~(p <=> q) <=> r) <=> ~(p <=> ~(q <=> r))`));
72004
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   525
val _ = pvFm (nnf (F`~((((p <=> q) <=> r) /\ (q <=> r)) ==> p)`));
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   526
val _ = pvFm (nnf (F`p <=> q`));
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   527
val _ = pvFm (nnf (F`p <=> q <=> r`));
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   528
val _ = pvFm (nnf (F`p <=> q <=> r <=> s`));
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
   529
val _ = pvFm (nnf (F`p <=> q <=> r <=> s <=> t`));
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   530
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   531
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   532
val () = SAY "Conjunctive normal form";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   533
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   534
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   535
local
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   536
  fun clauseToFormula cl =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   537
      Formula.listMkDisj (LiteralSet.transform Literal.toFormula cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   538
in
42102
fcfd07f122d4 new version of Metis 2.3 (29 Dec. 2010)
blanchet
parents: 39502
diff changeset
   539
  fun clausesToFormula cls = Formula.listMkConj (List.map clauseToFormula cls);
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   540
end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   541
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   542
val cnf' = pvFm o clausesToFormula o Normalize.cnf o F;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   543
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   544
val cnf = pvFm o clausesToFormula o Normalize.cnf o
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   545
          Formula.Not o Formula.generalize o F;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   546
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   547
val _ = cnf `p \/ ~p`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   548
val _ = cnf `~((p /\ (q \/ r /\ s)) /\ (~p \/ ~q \/ ~s))`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   549
val _ = cnf `~((p /\ (q \/ r /\ s)) /\ (~p \/ ~q \/ ~s) /\ (p \/ ~p))`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   550
val _ = cnf `~(~(p <=> q) <=> r) <=> ~(p <=> ~(q <=> r))`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   551
val _ = cnf `((p <=> q) <=> r) <=> (p <=> (q <=> r))`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   552
val _ = cnf `~(!x. ?y. x < y ==> !v. ?w. x * v < y * w)`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   553
val _ = cnf `~(!x. P x ==> (?y z. Q y \/ ~(?z. P z /\ Q z)))`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   554
val _ = cnf `~(?x y. x + y = 2)`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   555
val _ = cnf' `(!x. p x) \/ (!y. r $x y)`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   556
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   557
val _ = cnf
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   558
  `(!x. P x ==> (!x. Q x)) /\ ((!x. Q x \/ R x) ==> (?x. Q x /\ R x)) /\
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   559
   ((?x. R x) ==> (!x. L x ==> M x)) ==> (!x. P x /\ L x ==> M x)`;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   560
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   561
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   562
val () = SAY "Finite models";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   563
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   564
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   565
fun checkModelClause M cl =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   566
    let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   567
      val randomSamples = 100
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   568
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   569
      fun addRandomSample {T,F} =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   570
          let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   571
            val {T = T', F = F'} = Model.checkClause {maxChecks = SOME 1} M cl
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   572
          in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   573
            {T = T + T', F = F + F'}
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   574
          end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   575
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   576
      val {T,F} = funpow randomSamples addRandomSample {T = 0, F = 0}
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   577
      val rx = Real.fromInt T / Real.fromInt (T + F)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   578
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   579
      val {T,F} = Model.checkClause {maxChecks = NONE} M cl
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   580
      val ry = Real.fromInt T / Real.fromInt (T + F)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   581
    in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   582
      [Formula.toString (LiteralSet.disjoin cl),
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   583
       " | random sampling = " ^ percentToString rx,
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   584
       " | exhaustive = " ^ percentToString ry]
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   585
    end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   586
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   587
local
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   588
  val format =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   589
      [{leftAlign = true, padChar = #" "},
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   590
       {leftAlign = true, padChar = #" "},
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   591
       {leftAlign = true, padChar = #" "}];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   592
in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   593
  fun checkModel M cls =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   594
      let
42102
fcfd07f122d4 new version of Metis 2.3 (29 Dec. 2010)
blanchet
parents: 39502
diff changeset
   595
        val table = List.map (checkModelClause M) cls
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   596
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   597
        val rows = alignTable format table
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   598
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
   599
        val () = TextIO.print (join "\n" rows ^ "\n\n")
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   600
      in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   601
        ()
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   602
      end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   603
end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   604
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   605
fun perturbModel M cls n =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   606
    let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   607
      val N = {size = Model.size M}
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   608
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   609
      fun perturbClause (fv,cl) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   610
          let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   611
            val V = Model.randomValuation N fv
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   612
          in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   613
            if Model.interpretClause M V cl then ()
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   614
            else Model.perturbClause M V cl
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   615
          end
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   616
42102
fcfd07f122d4 new version of Metis 2.3 (29 Dec. 2010)
blanchet
parents: 39502
diff changeset
   617
      val cls = List.map (fn cl => (LiteralSet.freeVars cl, cl)) cls
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   618
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   619
      fun perturbClauses () = app perturbClause cls
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   620
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   621
      val () = funpow n perturbClauses ()
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   622
    in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   623
      M
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   624
    end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   625
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   626
val groupAxioms =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   627
    [LS[`0 + $x = $x`],
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   628
     LS[`~$x + $x = 0`],
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   629
     LS[`$x + $y + $z = $x + ($y + $z)`]];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   630
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   631
val groupThms =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   632
    [LS[`$x + 0 = $x`],
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   633
     LS[`$x + ~$x = 0`],
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   634
     LS[`~~$x = $x`]];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   635
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   636
fun newM fixed = Model.new {size = 8, fixed = fixed};
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   637
val M = pvM (newM Model.basicFixed);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   638
val () = checkModel M (groupAxioms @ groupThms);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   639
val M = pvM (perturbModel M groupAxioms 1000);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   640
val () = checkModel M (groupAxioms @ groupThms);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   641
val M = pvM (newM (Model.unionFixed Model.modularFixed Model.basicFixed));
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   642
val () = checkModel M (groupAxioms @ groupThms);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   643
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   644
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   645
val () = SAY "Checking the standard model";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   646
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   647
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   648
fun ppPercentClause (r,cl) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   649
    let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   650
      val ind = 6
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   651
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   652
      val p = percentToString r
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   653
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   654
      val fm = LiteralSet.disjoin cl
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   655
    in
45778
df6e210fb44c updated Metis to 20110926 version
blanchet
parents: 42102
diff changeset
   656
      Print.consistentBlock ind
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
   657
        [Print.ppString p,
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
   658
         Print.ppString (nChars #" " (ind - size p)),
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   659
         Formula.pp fm]
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   660
    end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   661
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   662
val standardModel = Model.new Model.default;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   663
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   664
fun checkStandardModelClause cl =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   665
    let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   666
      val {T,F} = Model.checkClause {maxChecks = SOME 1000} standardModel cl
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   667
      val r = Real.fromInt T / Real.fromInt (T + F)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   668
    in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   669
      (r,cl)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   670
    end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   671
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   672
val pvPCl = printval ppPercentClause
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   673
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   674
(* Equality *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   675
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   676
val cl = LS[`$x = $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   677
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   678
val cl = LS[`~($x = $y)`,`$y = $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   679
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   680
val cl = LS[`~($x = $y)`,`~($y = $z)`,`$x = $z`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   681
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   682
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   683
(* Projections *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   684
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   685
val cl = LS[`project1 $x1 = $x1`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   686
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   687
val cl = LS[`project1 $x1 $x2 = $x1`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   688
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   689
val cl = LS[`project2 $x1 $x2 = $x2`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   690
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   691
val cl = LS[`project1 $x1 $x2 $x3 = $x1`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   692
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   693
val cl = LS[`project2 $x1 $x2 $x3 = $x2`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   694
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   695
val cl = LS[`project3 $x1 $x2 $x3 = $x3`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   696
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   697
val cl = LS[`project1 $x1 $x2 $x3 $x4 = $x1`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   698
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   699
val cl = LS[`project2 $x1 $x2 $x3 $x4 = $x2`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   700
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   701
val cl = LS[`project3 $x1 $x2 $x3 $x4 = $x3`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   702
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   703
val cl = LS[`project4 $x1 $x2 $x3 $x4 = $x4`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   704
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   705
val cl = LS[`project1 $x1 $x2 $x3 $x4 $x5 = $x1`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   706
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   707
val cl = LS[`project2 $x1 $x2 $x3 $x4 $x5 = $x2`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   708
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   709
val cl = LS[`project3 $x1 $x2 $x3 $x4 $x5 = $x3`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   710
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   711
val cl = LS[`project4 $x1 $x2 $x3 $x4 $x5 = $x4`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   712
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   713
val cl = LS[`project5 $x1 $x2 $x3 $x4 $x5 = $x5`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   714
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   715
val cl = LS[`project1 $x1 $x2 $x3 $x4 $x5 $x6 = $x1`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   716
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   717
val cl = LS[`project2 $x1 $x2 $x3 $x4 $x5 $x6 = $x2`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   718
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   719
val cl = LS[`project3 $x1 $x2 $x3 $x4 $x5 $x6 = $x3`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   720
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   721
val cl = LS[`project4 $x1 $x2 $x3 $x4 $x5 $x6 = $x4`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   722
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   723
val cl = LS[`project5 $x1 $x2 $x3 $x4 $x5 $x6 = $x5`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   724
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   725
val cl = LS[`project6 $x1 $x2 $x3 $x4 $x5 $x6 = $x6`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   726
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   727
val cl = LS[`project1 $x1 $x2 $x3 $x4 $x5 $x6 $x7 = $x1`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   728
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   729
val cl = LS[`project2 $x1 $x2 $x3 $x4 $x5 $x6 $x7 = $x2`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   730
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   731
val cl = LS[`project3 $x1 $x2 $x3 $x4 $x5 $x6 $x7 = $x3`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   732
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   733
val cl = LS[`project4 $x1 $x2 $x3 $x4 $x5 $x6 $x7 = $x4`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   734
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   735
val cl = LS[`project5 $x1 $x2 $x3 $x4 $x5 $x6 $x7 = $x5`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   736
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   737
val cl = LS[`project6 $x1 $x2 $x3 $x4 $x5 $x6 $x7 = $x6`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   738
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   739
val cl = LS[`project7 $x1 $x2 $x3 $x4 $x5 $x6 $x7 = $x7`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   740
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   741
val cl = LS[`project1 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 = $x1`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   742
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   743
val cl = LS[`project2 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 = $x2`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   744
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   745
val cl = LS[`project3 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 = $x3`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   746
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   747
val cl = LS[`project4 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 = $x4`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   748
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   749
val cl = LS[`project5 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 = $x5`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   750
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   751
val cl = LS[`project6 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 = $x6`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   752
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   753
val cl = LS[`project7 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 = $x7`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   754
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   755
val cl = LS[`project8 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 = $x8`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   756
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   757
val cl = LS[`project1 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x1`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   758
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   759
val cl = LS[`project2 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x2`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   760
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   761
val cl = LS[`project3 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x3`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   762
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   763
val cl = LS[`project4 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x4`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   764
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   765
val cl = LS[`project5 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x5`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   766
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   767
val cl = LS[`project6 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x6`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   768
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   769
val cl = LS[`project7 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x7`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   770
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   771
val cl = LS[`project8 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x8`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   772
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   773
val cl = LS[`project9 $x1 $x2 $x3 $x4 $x5 $x6 $x7 $x8 $x9 = $x9`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   774
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   775
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   776
(* Arithmetic *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   777
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   778
(* Zero *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   779
val cl = LS[`~isZero $x`,`$x = 0`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   780
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   781
val cl = LS[`isZero $x`,`~($x = 0)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   782
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   783
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   784
(* Positive numerals *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   785
val cl = LS[`0 + 1 = 1`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   786
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   787
val cl = LS[`1 + 1 = 2`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   788
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   789
val cl = LS[`2 + 1 = 3`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   790
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   791
val cl = LS[`3 + 1 = 4`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   792
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   793
val cl = LS[`4 + 1 = 5`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   794
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   795
val cl = LS[`5 + 1 = 6`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   796
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   797
val cl = LS[`6 + 1 = 7`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   798
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   799
val cl = LS[`7 + 1 = 8`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   800
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   801
val cl = LS[`8 + 1 = 9`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   802
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   803
val cl = LS[`9 + 1 = 10`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   804
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   805
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   806
(* Negative numerals *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   807
val cl = LS[`~1 = negative1`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   808
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   809
val cl = LS[`~2 = negative2`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   810
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   811
val cl = LS[`~3 = negative3`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   812
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   813
val cl = LS[`~4 = negative4`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   814
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   815
val cl = LS[`~5 = negative5`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   816
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   817
val cl = LS[`~6 = negative6`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   818
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   819
val cl = LS[`~7 = negative7`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   820
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   821
val cl = LS[`~8 = negative8`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   822
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   823
val cl = LS[`~9 = negative9`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   824
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   825
val cl = LS[`~10 = negative10`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   826
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   827
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   828
(* Addition *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   829
val cl = LS[`0 + $x = $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   830
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   831
val cl = LS[`$x + $y = $y + $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   832
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   833
val cl = LS[`$x + ($y + $z) = ($x + $y) + $z`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   834
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   835
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   836
(* Negation *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   837
val cl = LS[`~$x + $x = 0`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   838
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   839
val cl = LS[`~~$x = $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   840
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   841
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   842
(* Subtraction *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   843
val cl = LS[`$x - $y = $x + ~$y`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   844
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   845
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   846
(* Successor *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   847
val cl = LS[`suc $x = $x + 1`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   848
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   849
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   850
(* Predecessor *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   851
val cl = LS[`pre $x = $x - 1`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   852
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   853
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   854
(* Ordering *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   855
val cl = LS[`$x <= $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   856
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   857
val cl = LS[`~($x <= $y)`,`~($y <= $z)`,`$x <= $z`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   858
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   859
val cl = LS[`~($x <= $y)`,`~($y <= $x)`,`$x = $y`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   860
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   861
val cl = LS[`0 <= $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   862
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   863
val cl = LS[`~($x >= $y)`,`$y <= $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   864
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   865
val cl = LS[`$x >= $y`,`~($y <= $x)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   866
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   867
val cl = LS[`$x > $y`,`$x <= $y`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   868
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   869
val cl = LS[`~($x > $y)`,`~($x <= $y)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   870
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   871
val cl = LS[`$x < $y`,`$y <= $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   872
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   873
val cl = LS[`~($x < $y)`,`~($y <= $x)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   874
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   875
val cl = LS[`$x = 0`,`~($x <= $y)`,`~$y <= ~$x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   876
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   877
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   878
(* Multiplication *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   879
val cl = LS[`1 * $x = $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   880
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   881
val cl = LS[`0 * $x = 0`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   882
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   883
val cl = LS[`$x * $y = $y * $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   884
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   885
val cl = LS[`$x * ($y * $z) = ($x * $y) * $z`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   886
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   887
val cl = LS[`$x * ($y + $z) = ($x * $y) + ($x * $z)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   888
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   889
val cl = LS[`$x * ~$y = ~($x * $y)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   890
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   891
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   892
(* Division *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   893
val cl = LS[`$y = 0`,`$x mod $y < $y`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   894
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   895
val cl = LS[`$y * ($x div $y) + $x mod $y = $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   896
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   897
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   898
(* Exponentiation *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   899
val cl = LS[`exp $x 0 = 1`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   900
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   901
val cl = LS[`$y = 0`,`exp $x $y = $x * exp $x (pre $y)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   902
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   903
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   904
(* Divides *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   905
val cl = LS[`divides $x $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   906
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   907
val cl = LS[`~(divides $x $y)`,`~(divides $y $z)`,`divides $x $z`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   908
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   909
val cl = LS[`~(divides $x $y)`,`~(divides $y $x)`,`$x = $y`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   910
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   911
val cl = LS[`divides 1 $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   912
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   913
val cl = LS[`divides $x 0`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   914
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   915
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   916
(* Even and odd *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   917
val cl = LS[`even 0`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   918
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   919
val cl = LS[`$x = 0`,`~(even (pre $x))`,`odd $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   920
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   921
val cl = LS[`$x = 0`,`~(odd (pre $x))`,`even $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   922
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   923
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   924
(* Sets *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   925
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   926
(* The empty set *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   927
val cl = LS[`~member $x empty`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   928
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   929
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   930
(* The universal set *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   931
val cl = LS[`member $x universe`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   932
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   933
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   934
(* Complement *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   935
val cl = LS[`member $x $y`,`member $x (complement $y)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   936
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   937
val cl = LS[`~(member $x $y)`,`~member $x (complement $y)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   938
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   939
val cl = LS[`complement (complement $x) = $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   940
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   941
val cl = LS[`complement empty = universe`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   942
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   943
val cl = LS[`complement universe = empty`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   944
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   945
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   946
(* The subset relation *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   947
val cl = LS[`subset $x $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   948
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   949
val cl = LS[`~subset $x $y`,`~subset $y $z`,`subset $x $z`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   950
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   951
val cl = LS[`~subset $x $y`,`~subset $y $x`,`$x = $y`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   952
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   953
val cl = LS[`subset empty $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   954
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   955
val cl = LS[`subset $x universe`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   956
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   957
val cl = LS[`~subset $x $y`,`subset (complement $y) (complement $x)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   958
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   959
val cl = LS[`~member $x $y`,`~subset $y $z`,`member $x $z`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   960
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   961
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   962
(* Union *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   963
val cl = LS[`union $x $y = union $y $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   964
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   965
val cl = LS[`union $x (union $y $z) = union (union $x $y) $z`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   966
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   967
val cl = LS[`union empty $x = $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   968
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   969
val cl = LS[`union universe $x = universe`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   970
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   971
val cl = LS[`subset $x (union $x $y)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   972
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   973
val cl = LS[`~member $x (union $y $z)`,`member $x $y`,`member $x $z`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   974
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   975
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   976
(* Intersection *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   977
val cl = LS[`intersect $x $y =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   978
             complement (union (complement $x) (complement $y))`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   979
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   980
val cl = LS[`subset (intersect $x $y) $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   981
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   982
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   983
(* Difference *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   984
val cl = LS[`difference $x $y = intersect $x (complement $y)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   985
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   986
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   987
(* Symmetric difference *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   988
val cl = LS[`symmetricDifference $x $y =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   989
             union (difference $x $y) (difference $y $x)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   990
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   991
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   992
(* Insert *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   993
val cl = LS[`member $x (insert $x $y)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   994
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   995
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   996
(* Singleton *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   997
val cl = LS[`singleton $x = (insert $x empty)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   998
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
   999
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1000
(* Cardinality *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1001
val cl = LS[`card empty = 0`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1002
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1003
val cl = LS[`member $x $y`,`card (insert $x $y) = suc (card $y)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1004
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1005
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1006
(* Lists *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1007
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1008
(* Nil *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1009
val cl = LS[`null nil`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1010
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1011
val cl = LS[`~null $x`, `$x = nil`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1012
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1013
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1014
(* Cons *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1015
val cl = LS[`~(nil = $x :: $y)`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1016
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1017
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1018
(* Append *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1019
val cl = LS[`$x @ ($y @ $z) = ($x @ $y) @ $z`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1020
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1021
val cl = LS[`nil @ $x = $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1022
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1023
val cl = LS[`$x @ nil = $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1024
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1025
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1026
(* Length *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1027
val cl = LS[`length nil = 0`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1028
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1029
val cl = LS[`length ($x :: $y) >= length $y`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1030
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1031
val cl = LS[`length ($x @ $y) >= length $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1032
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1033
val cl = LS[`length ($x @ $y) >= length $y`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1034
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1035
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1036
(* Tail *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1037
val cl = LS[`null $x`,`suc (length (tail $x)) = length $x`];
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1038
val _ = pvPCl (checkStandardModelClause cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1039
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1040
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1041
val () = SAY "Clauses";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1042
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1043
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1044
val cl = pvCl (CL[`P $x`,`P $y`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1045
val _ = pvLits (Clause.largestLiterals cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1046
val _ = pvCls (Clause.factor cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1047
val cl = pvCl (CL[`P $x`,`~P (f $x)`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1048
val _ = pvLits (Clause.largestLiterals cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1049
val cl = pvCl (CL[`$x = $y`,`f $y = f $x`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1050
val _ = pvLits (Clause.largestLiterals cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1051
val cl = pvCl (CL[`$x = f $y`,`f $x = $y`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1052
val _ = pvLits (Clause.largestLiterals cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1053
val cl = pvCl (CL[`s = a`,`s = b`,`h b c`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1054
val _ = pvLits (Clause.largestLiterals cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1055
val cl = pvCl (CL[`a = a`,`a = b`,`h b c`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1056
val _ = pvLits (Clause.largestLiterals cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1057
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1058
(* Test cases contributed by Larry Paulson *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1059
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1060
local
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1061
  val lnFnName = Name.fromString "ln"
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1062
  and expFnName = Name.fromString "exp"
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1063
  and divFnName = Name.fromString "/"
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1064
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1065
  val leRelName = Name.fromString "<";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1066
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1067
  fun weight na =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1068
      case na of
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1069
        (n,1) =>
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1070
        if Name.equal n lnFnName then 500
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1071
        else if Name.equal n expFnName then 500
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1072
        else 1
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1073
      | (n,2) =>
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1074
        if Name.equal n divFnName then 50
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1075
        else if Name.equal n leRelName then 20
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1076
        else 1
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1077
      | _ => 1;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1078
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1079
  val ordering =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1080
      {weight = weight, precedence = #precedence KnuthBendixOrder.default};
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1081
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1082
  val clauseParameters =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1083
      {ordering = ordering,
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1084
       orderLiterals = Clause.UnsignedLiteralOrder,
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1085
       orderTerms = true};
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1086
in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1087
  val LcpCL = mkCl clauseParameters o AX;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1088
end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1089
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1090
val cl = pvCl (LcpCL[`~($y <= (2 + (2 * $x + pow $x 2)) / 2)`, `~(0 <= $x)`,
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1091
                     `$y <= exp $x`]);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1092
val _ = pvLits (Clause.largestLiterals cl);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1093
72004
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1094
(* Bug discovered by Michael Farber *)
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1095
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1096
local
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1097
  fun activeFactor th =
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1098
      let
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1099
        val (_,{axioms,conjecture}) =
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1100
            Active.new Active.default {axioms = [], conjecture = [th]}
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1101
      in
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1102
        List.map Clause.thm (axioms @ conjecture)
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1103
      end;
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1104
in
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1105
  val th = pvThm (AX[`c4 (c5 (c6 c7 c8) $y) $z = c3`,
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1106
                     `c4 (c5 (c6 c7 c8) $t) $u = c3`]);
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1107
  val _ = pvThms (activeFactor th);
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1108
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1109
  val th = pvThm (AX[`~(c4 (c5 (c6 c7 c8) c28) c29 = c4 (c5 (c6 c7 c8) c28) $x)`,
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1110
                     `c4 (c5 (c6 c7 c8) $y) $z = c3`,
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1111
                     `c4 (c5 (c6 c7 c8) $t) $u = c3`]);
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1112
  val _ = pvThms (activeFactor th);
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1113
end;
913162a47d9f Update Metis to 2.4
desharna
parents: 45778
diff changeset
  1114
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1115
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1116
val () = SAY "Syntax checking the problem sets";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1117
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1118
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1119
local
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1120
  fun same n = raise Fail ("Two goals called " ^ n);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1121
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1122
  fun dup n n' =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1123
      raise Fail ("Goal " ^ n' ^ " is probable duplicate of " ^ n);
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1124
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1125
  fun quot fm =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1126
      let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1127
        fun f (v,s) = Subst.insert s (v,V"_")
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1128
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1129
        val sub = NameSet.foldl f Subst.empty (Formula.freeVars fm)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1130
      in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1131
        Formula.subst sub fm
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1132
      end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1133
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1134
  val quot_clauses =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1135
      Formula.listMkConj o sort Formula.compare o
42102
fcfd07f122d4 new version of Metis 2.3 (29 Dec. 2010)
blanchet
parents: 39502
diff changeset
  1136
      List.map (quot o snd o Formula.stripForall) o Formula.stripConj;
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1137
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1138
  fun quotient (Formula.Imp (a, Formula.Imp (b, Formula.False))) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1139
      Formula.Imp (quot_clauses a, Formula.Imp (quot_clauses b, Formula.False))
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1140
    | quotient fm = fm;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1141
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1142
  fun check ({name,goal,...}, acc) =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1143
    let
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1144
      val g = prep goal
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1145
      val p =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1146
          Formula.fromString g
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1147
          handle Parse.NoParse =>
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1148
                 raise Error ("failed to parse problem " ^ name)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1149
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1150
      val () =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1151
        case List.find (fn (n,_) => n = name) acc of NONE => ()
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1152
        | SOME _ => same name
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1153
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1154
      val () =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1155
        case List.find (fn (_,x) => Formula.equal x p) acc of NONE => ()
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1156
        | SOME (n,_) => dup n name
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1157
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1158
      val _ =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1159
        test_fun equal I g (mini_print (!Print.lineLength) p)
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
  1160
        handle e =>
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
  1161
          (TextIO.print ("Error in problem " ^ name ^ "\n\n");
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
  1162
           raise e)
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1163
    in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1164
      (name,p) :: acc
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1165
    end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1166
in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1167
  fun check_syntax (p : problem list) =
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
  1168
      let
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
  1169
        val _ = List.foldl check [] p
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
  1170
      in
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
  1171
        TextIO.print "ok\n\n"
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
  1172
      end;
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1173
end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1174
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1175
val () = check_syntax problems;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1176
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1177
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1178
val () = SAY "Parsing TPTP problems";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1179
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1180
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1181
fun tptp f =
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1182
    let
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
  1183
      val () = TextIO.print ("parsing " ^ f ^ "... ")
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1184
      val filename = "tptp/" ^ f ^ ".tptp"
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1185
      val mapping = Tptp.defaultMapping
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1186
      val goal = Tptp.goal (Tptp.read {filename = filename, mapping = mapping})
39443
e330437cd22a copied the unmodified official Metis 2.3 (15 Sept. 2010) sources into Isabelle
blanchet
parents: 39353
diff changeset
  1187
      val () = TextIO.print "ok\n"
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1188
    in
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1189
      pvFm goal
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1190
    end;
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1191
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1192
val _ = tptp "PUZ001-1";
42102
fcfd07f122d4 new version of Metis 2.3 (29 Dec. 2010)
blanchet
parents: 39502
diff changeset
  1193
val _ = tptp "NO_FORMULAS";
fcfd07f122d4 new version of Metis 2.3 (29 Dec. 2010)
blanchet
parents: 39502
diff changeset
  1194
val _ = tptp "SEPARATED_COMMENTS";
39348
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1195
val _ = tptp "NUMBERED_FORMULAS";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1196
val _ = tptp "DEFINED_TERMS";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1197
val _ = tptp "SYSTEM_TERMS";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1198
val _ = tptp "QUOTED_TERMS";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1199
val _ = tptp "QUOTED_TERMS_IDENTITY";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1200
val _ = tptp "QUOTED_TERMS_DIFFERENT";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1201
val _ = tptp "QUOTED_TERMS_SPECIAL";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1202
val _ = tptp "RENAMING_VARIABLES";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1203
val _ = tptp "MIXED_PROBLEM";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1204
val _ = tptp "BLOCK_COMMENTS";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1205
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1206
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1207
val () = SAY "The TPTP finite model";
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1208
(* ------------------------------------------------------------------------- *)
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1209
6f9c9899f99f new version of the Metis files
blanchet
parents:
diff changeset
  1210
val _ = printval (Tptp.ppFixedMap Tptp.defaultMapping) Tptp.defaultFixedMap;