| author | wenzelm | 
| Mon, 01 May 2006 17:05:09 +0200 | |
| changeset 19521 | cfdab6a91332 | 
| parent 19447 | d4f3f8f9c0a5 | 
| child 19642 | ea7162f84677 | 
| permissions | -rw-r--r-- | 
| 15347 | 1  | 
(* Author: Jia Meng, Cambridge University Computer Laboratory  | 
2  | 
ID: $Id$  | 
|
3  | 
Copyright 2004 University of Cambridge  | 
|
4  | 
||
5  | 
ML data structure for storing/printing FOL clauses and arity clauses.  | 
|
6  | 
Typed equality is treated differently.  | 
|
7  | 
*)  | 
|
8  | 
||
9  | 
signature RES_CLAUSE =  | 
|
10  | 
sig  | 
|
| 18868 | 11  | 
exception CLAUSE of string * term  | 
12  | 
type clause and arityClause and classrelClause  | 
|
13  | 
type fol_type  | 
|
14  | 
type typ_var  | 
|
15  | 
type type_literal  | 
|
16  | 
val add_typs_aux : (typ_var * string list) list -> type_literal list * type_literal list  | 
|
17  | 
val arity_clause_thy: theory -> arityClause list  | 
|
18  | 
val ascii_of : string -> string  | 
|
19  | 
val bracket_pack : string list -> string  | 
|
20  | 
  val check_var_pairs: ''a * ''b -> (''a * ''b) list -> int
 | 
|
21  | 
val classrel_clauses_thy: theory -> classrelClause list  | 
|
22  | 
val clause_eq : clause * clause -> bool  | 
|
23  | 
val clause_prefix : string  | 
|
24  | 
val clause2tptp : clause -> string * string list  | 
|
25  | 
val const_prefix : string  | 
|
26  | 
val dfg_write_file: thm list -> string -> (clause list * classrelClause list * arityClause list) -> unit  | 
|
27  | 
val fixed_var_prefix : string  | 
|
28  | 
val gen_tptp_cls : int * string * string * string -> string  | 
|
29  | 
val gen_tptp_type_cls : int * string * string * string * int -> string  | 
|
30  | 
val get_axiomName : clause -> string  | 
|
31  | 
val hash_clause : clause -> int  | 
|
32  | 
val init : theory -> unit  | 
|
33  | 
val isMeta : string -> bool  | 
|
34  | 
val isTaut : clause -> bool  | 
|
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
35  | 
val keep_types : bool ref  | 
| 18868 | 36  | 
  val list_ord : ('a * 'b -> order) -> 'a list * 'b list -> order
 | 
| 
19443
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
37  | 
val make_axiom_clause : thm -> string * int -> clause option  | 
| 
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
38  | 
val make_conjecture_clauses : thm list -> clause list  | 
| 18868 | 39  | 
val make_fixed_const : string -> string  | 
40  | 
val make_fixed_type_const : string -> string  | 
|
41  | 
val make_fixed_type_var : string -> string  | 
|
42  | 
val make_fixed_var : string -> string  | 
|
43  | 
val make_schematic_type_var : string * int -> string  | 
|
44  | 
val make_schematic_var : string * int -> string  | 
|
45  | 
val make_type_class : string -> string  | 
|
46  | 
val mk_fol_type: string * string * fol_type list -> fol_type  | 
|
47  | 
val mk_typ_var_sort : Term.typ -> typ_var * sort  | 
|
48  | 
val paren_pack : string list -> string  | 
|
49  | 
val schematic_var_prefix : string  | 
|
50  | 
val special_equal : bool ref  | 
|
51  | 
val string_of_fol_type : fol_type -> string  | 
|
52  | 
val tconst_prefix : string  | 
|
53  | 
val tfree_prefix : string  | 
|
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
54  | 
val tptp_arity_clause : arityClause -> string  | 
| 
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
55  | 
val tptp_classrelClause : classrelClause -> string  | 
| 18868 | 56  | 
val tptp_of_typeLit : type_literal -> string  | 
| 18863 | 57  | 
val tptp_tfree_clause : string -> string  | 
| 
19443
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
58  | 
val tptp_write_file: thm list -> string -> ((thm * (string * int)) list * classrelClause list * arityClause list) -> unit  | 
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
59  | 
val tvar_prefix : string  | 
| 18868 | 60  | 
val types_eq: fol_type list * fol_type list -> (string*string) list * (string*string) list -> bool * ((string*string) list * (string*string) list)  | 
61  | 
val types_ord : fol_type list * fol_type list -> order  | 
|
| 
17908
 
ac97527724ba
More functions are added to the signature of ResClause
 
mengj 
parents: 
17888 
diff
changeset
 | 
62  | 
val union_all : ''a list list -> ''a list  | 
| 18863 | 63  | 
val writeln_strs: TextIO.outstream -> TextIO.vector list -> unit  | 
| 15347 | 64  | 
end;  | 
65  | 
||
| 
18420
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
66  | 
structure ResClause : RES_CLAUSE =  | 
| 15347 | 67  | 
struct  | 
68  | 
||
69  | 
(* Added for typed equality *)  | 
|
70  | 
val special_equal = ref false; (* by default,equality does not carry type information *)  | 
|
71  | 
val eq_typ_wrapper = "typeinfo"; (* default string *)  | 
|
72  | 
||
73  | 
||
74  | 
val schematic_var_prefix = "V_";  | 
|
75  | 
val fixed_var_prefix = "v_";  | 
|
76  | 
||
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
77  | 
val tvar_prefix = "T_";  | 
| 
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
78  | 
val tfree_prefix = "t_";  | 
| 15347 | 79  | 
|
80  | 
val clause_prefix = "cls_";  | 
|
| 
17525
 
ae5bb6001afb
tidying, and support for axclass/classrel clauses
 
paulson 
parents: 
17422 
diff
changeset
 | 
81  | 
val arclause_prefix = "clsarity_"  | 
| 
 
ae5bb6001afb
tidying, and support for axclass/classrel clauses
 
paulson 
parents: 
17422 
diff
changeset
 | 
82  | 
val clrelclause_prefix = "clsrel_";  | 
| 15347 | 83  | 
|
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
84  | 
val const_prefix = "c_";  | 
| 
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
85  | 
val tconst_prefix = "tc_";  | 
| 
16199
 
ee95ab217fee
no longer emits literals for type class HOL.type; also minor tidying
 
paulson 
parents: 
16039 
diff
changeset
 | 
86  | 
val class_prefix = "class_";  | 
| 15347 | 87  | 
|
| 17775 | 88  | 
fun union_all xss = foldl (op union) [] xss;  | 
89  | 
||
90  | 
(*Provide readable names for the more common symbolic functions*)  | 
|
| 15347 | 91  | 
val const_trans_table =  | 
92  | 
      Symtab.make [("op =", "equal"),
 | 
|
| 19277 | 93  | 
	  	   ("Orderings.less_eq", "lessequals"),
 | 
94  | 
		   ("Orderings.less", "less"),
 | 
|
| 15347 | 95  | 
		   ("op &", "and"),
 | 
96  | 
		   ("op |", "or"),
 | 
|
| 
19233
 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 
haftmann 
parents: 
19207 
diff
changeset
 | 
97  | 
		   ("HOL.plus", "plus"),
 | 
| 
 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 
haftmann 
parents: 
19207 
diff
changeset
 | 
98  | 
		   ("HOL.minus", "minus"),
 | 
| 
 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 
haftmann 
parents: 
19207 
diff
changeset
 | 
99  | 
		   ("HOL.times", "times"),
 | 
| 18676 | 100  | 
		   ("Divides.op div", "div"),
 | 
101  | 
		   ("HOL.divide", "divide"),
 | 
|
| 15347 | 102  | 
		   ("op -->", "implies"),
 | 
| 17375 | 103  | 
		   ("{}", "emptyset"),
 | 
| 15347 | 104  | 
		   ("op :", "in"),
 | 
105  | 
		   ("op Un", "union"),
 | 
|
| 18390 | 106  | 
		   ("op Int", "inter"),
 | 
107  | 
		   ("List.op @", "append")];
 | 
|
| 15347 | 108  | 
|
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
109  | 
val type_const_trans_table =  | 
| 
18411
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
110  | 
      Symtab.make [("*", "prod"),
 | 
| 
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
111  | 
	  	   ("+", "sum"),
 | 
| 
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
112  | 
		   ("~=>", "map")];
 | 
| 15347 | 113  | 
|
| 15610 | 114  | 
(*Escaping of special characters.  | 
115  | 
Alphanumeric characters are left unchanged.  | 
|
116  | 
The character _ goes to __  | 
|
117  | 
Characters in the range ASCII space to / go to _A to _P, respectively.  | 
|
118  | 
Other printing characters go to _NNN where NNN is the decimal ASCII code.*)  | 
|
119  | 
local  | 
|
120  | 
||
121  | 
val A_minus_space = Char.ord #"A" - Char.ord #" ";  | 
|
122  | 
||
| 15347 | 123  | 
fun ascii_of_c c =  | 
| 15610 | 124  | 
if Char.isAlphaNum c then String.str c  | 
125  | 
else if c = #"_" then "__"  | 
|
126  | 
else if #" " <= c andalso c <= #"/"  | 
|
127  | 
then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))  | 
|
128  | 
  else if Char.isPrint c then ("_" ^ Int.toString (Char.ord c))
 | 
|
129  | 
else ""  | 
|
| 15347 | 130  | 
|
| 15610 | 131  | 
in  | 
132  | 
||
133  | 
val ascii_of = String.translate ascii_of_c;  | 
|
134  | 
||
135  | 
end;  | 
|
| 15347 | 136  | 
|
| 
17525
 
ae5bb6001afb
tidying, and support for axclass/classrel clauses
 
paulson 
parents: 
17422 
diff
changeset
 | 
137  | 
(* convert a list of strings into one single string; surrounded by brackets *)  | 
| 
18218
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
138  | 
fun paren_pack [] = "" (*empty argument list*)  | 
| 
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
139  | 
  | paren_pack strings = "(" ^ commas strings ^ ")";
 | 
| 
17525
 
ae5bb6001afb
tidying, and support for axclass/classrel clauses
 
paulson 
parents: 
17422 
diff
changeset
 | 
140  | 
|
| 
 
ae5bb6001afb
tidying, and support for axclass/classrel clauses
 
paulson 
parents: 
17422 
diff
changeset
 | 
141  | 
fun bracket_pack strings = "[" ^ commas strings ^ "]";  | 
| 
 
ae5bb6001afb
tidying, and support for axclass/classrel clauses
 
paulson 
parents: 
17422 
diff
changeset
 | 
142  | 
|
| 
 
ae5bb6001afb
tidying, and support for axclass/classrel clauses
 
paulson 
parents: 
17422 
diff
changeset
 | 
143  | 
|
| 
16925
 
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
 
paulson 
parents: 
16903 
diff
changeset
 | 
144  | 
(*Remove the initial ' character from a type variable, if it is present*)  | 
| 
 
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
 
paulson 
parents: 
16903 
diff
changeset
 | 
145  | 
fun trim_type_var s =  | 
| 
 
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
 
paulson 
parents: 
16903 
diff
changeset
 | 
146  | 
if s <> "" andalso String.sub(s,0) = #"'" then String.extract(s,1,NONE)  | 
| 
 
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
 
paulson 
parents: 
16903 
diff
changeset
 | 
147  | 
  else error ("trim_type: Malformed type variable encountered: " ^ s);
 | 
| 
 
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
 
paulson 
parents: 
16903 
diff
changeset
 | 
148  | 
|
| 16903 | 149  | 
fun ascii_of_indexname (v,0) = ascii_of v  | 
| 
17525
 
ae5bb6001afb
tidying, and support for axclass/classrel clauses
 
paulson 
parents: 
17422 
diff
changeset
 | 
150  | 
| ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ Int.toString i;  | 
| 15347 | 151  | 
|
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
152  | 
fun make_schematic_var v = schematic_var_prefix ^ (ascii_of_indexname v);  | 
| 15347 | 153  | 
fun make_fixed_var x = fixed_var_prefix ^ (ascii_of x);  | 
154  | 
||
| 
16925
 
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
 
paulson 
parents: 
16903 
diff
changeset
 | 
155  | 
fun make_schematic_type_var (x,i) =  | 
| 
 
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
 
paulson 
parents: 
16903 
diff
changeset
 | 
156  | 
tvar_prefix ^ (ascii_of_indexname (trim_type_var x,i));  | 
| 
 
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
 
paulson 
parents: 
16903 
diff
changeset
 | 
157  | 
fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x));  | 
| 15347 | 158  | 
|
| 
18411
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
159  | 
fun lookup_const c =  | 
| 17412 | 160  | 
case Symtab.lookup const_trans_table c of  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
161  | 
SOME c' => c'  | 
| 
18411
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
162  | 
| NONE => ascii_of c;  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
163  | 
|
| 
18411
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
164  | 
fun lookup_type_const c =  | 
| 17412 | 165  | 
case Symtab.lookup type_const_trans_table c of  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
166  | 
SOME c' => c'  | 
| 
18411
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
167  | 
| NONE => ascii_of c;  | 
| 
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
168  | 
|
| 
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
169  | 
fun make_fixed_const "op =" = "equal" (*MUST BE "equal" because it's built-in to ATPs*)  | 
| 
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
170  | 
| make_fixed_const c = const_prefix ^ lookup_const c;  | 
| 
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
171  | 
|
| 
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
172  | 
fun make_fixed_type_const c = tconst_prefix ^ lookup_type_const c;  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
173  | 
|
| 17261 | 174  | 
fun make_type_class clas = class_prefix ^ ascii_of clas;  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
175  | 
|
| 
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
176  | 
|
| 18798 | 177  | 
(***** definitions and functions for FOL clauses, for conversion to TPTP or DFG format. *****)  | 
| 15347 | 178  | 
|
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
179  | 
val keep_types = ref true;  | 
| 15347 | 180  | 
|
181  | 
datatype kind = Axiom | Hypothesis | Conjecture;  | 
|
182  | 
fun name_of_kind Axiom = "axiom"  | 
|
183  | 
| name_of_kind Hypothesis = "hypothesis"  | 
|
184  | 
| name_of_kind Conjecture = "conjecture";  | 
|
185  | 
||
186  | 
type clause_id = int;  | 
|
187  | 
type axiom_name = string;  | 
|
188  | 
||
189  | 
||
190  | 
type polarity = bool;  | 
|
191  | 
||
| 
18869
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
192  | 
(* "tag" is used for vampire specific syntax FIXME REMOVE *)  | 
| 15347 | 193  | 
type tag = bool;  | 
194  | 
||
195  | 
||
196  | 
(**** Isabelle FOL clauses ****)  | 
|
197  | 
||
198  | 
val tagged = ref false;  | 
|
199  | 
||
200  | 
type pred_name = string;  | 
|
| 
18402
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
201  | 
|
| 
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
202  | 
datatype typ_var = FOLTVar of indexname | FOLTFree of string;  | 
| 
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
203  | 
|
| 18798 | 204  | 
(*FIXME: give the constructors more sensible names*)  | 
| 
18402
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
205  | 
datatype fol_type = AtomV of string  | 
| 
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
206  | 
| AtomF of string  | 
| 
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
207  | 
| Comp of string * fol_type list;  | 
| 
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
208  | 
|
| 
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
209  | 
fun string_of_fol_type (AtomV x) = x  | 
| 
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
210  | 
| string_of_fol_type (AtomF x) = x  | 
| 
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
211  | 
| string_of_fol_type (Comp(tcon,tps)) =  | 
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
212  | 
tcon ^ (paren_pack (map string_of_fol_type tps));  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
213  | 
|
| 
18439
 
4b517881ac7e
Added functions for fol_type; also put some functions to the signature, used by ResHolClause.
 
mengj 
parents: 
18420 
diff
changeset
 | 
214  | 
fun mk_fol_type ("Var",x,_) = AtomV(x)
 | 
| 
 
4b517881ac7e
Added functions for fol_type; also put some functions to the signature, used by ResHolClause.
 
mengj 
parents: 
18420 
diff
changeset
 | 
215  | 
  | mk_fol_type ("Fixed",x,_) = AtomF(x)
 | 
| 
 
4b517881ac7e
Added functions for fol_type; also put some functions to the signature, used by ResHolClause.
 
mengj 
parents: 
18420 
diff
changeset
 | 
216  | 
  | mk_fol_type ("Comp",con,args) = Comp(con,args)
 | 
| 
 
4b517881ac7e
Added functions for fol_type; also put some functions to the signature, used by ResHolClause.
 
mengj 
parents: 
18420 
diff
changeset
 | 
217  | 
|
| 
 
4b517881ac7e
Added functions for fol_type; also put some functions to the signature, used by ResHolClause.
 
mengj 
parents: 
18420 
diff
changeset
 | 
218  | 
|
| 18798 | 219  | 
(*First string is the type class; the second is a TVar or TFfree*)  | 
220  | 
datatype type_literal = LTVar of string * string | LTFree of string * string;  | 
|
| 15347 | 221  | 
|
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
222  | 
datatype fol_term = UVar of string * fol_type  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
223  | 
| Fun of string * fol_type list * fol_term list;  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
224  | 
datatype predicate = Predicate of pred_name * fol_type list * fol_term list;  | 
| 15347 | 225  | 
|
226  | 
datatype literal = Literal of polarity * predicate * tag;  | 
|
227  | 
||
| 17999 | 228  | 
fun mk_typ_var_sort (TFree(a,s)) = (FOLTFree a,s)  | 
229  | 
| mk_typ_var_sort (TVar(v,s)) = (FOLTVar v,s);  | 
|
230  | 
||
231  | 
||
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
232  | 
(*A clause has first-order literals and other, type-related literals*)  | 
| 15347 | 233  | 
datatype clause =  | 
234  | 
	 Clause of {clause_id: clause_id,
 | 
|
235  | 
axiom_name: axiom_name,  | 
|
| 19447 | 236  | 
th: thm,  | 
| 15347 | 237  | 
kind: kind,  | 
238  | 
literals: literal list,  | 
|
| 
18869
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
239  | 
types_sorts: (typ_var * sort) list};  | 
| 
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
240  | 
|
| 
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
241  | 
fun get_axiomName (Clause cls) = #axiom_name cls;  | 
| 15347 | 242  | 
|
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
243  | 
exception CLAUSE of string * term;  | 
| 15347 | 244  | 
|
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
245  | 
fun isFalse (Literal (pol,Predicate(pname,_,[]),_)) =  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
246  | 
(pol andalso pname = "c_False") orelse  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
247  | 
(not pol andalso pname = "c_True")  | 
| 
17317
 
3f12de2e2e6e
Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
 
paulson 
parents: 
17312 
diff
changeset
 | 
248  | 
| isFalse _ = false;  | 
| 
 
3f12de2e2e6e
Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
 
paulson 
parents: 
17312 
diff
changeset
 | 
249  | 
|
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
250  | 
fun isTrue (Literal (pol,Predicate(pname,_,[]),_)) =  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
251  | 
(pol andalso pname = "c_True") orelse  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
252  | 
(not pol andalso pname = "c_False")  | 
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
253  | 
| isTrue _ = false;  | 
| 
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
254  | 
|
| 
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
255  | 
fun isTaut (Clause {literals,...}) = exists isTrue literals;  
 | 
| 
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
256  | 
|
| 19447 | 257  | 
fun make_clause (clause_id, axiom_name, th, kind, literals, types_sorts) =  | 
| 
17317
 
3f12de2e2e6e
Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
 
paulson 
parents: 
17312 
diff
changeset
 | 
258  | 
if forall isFalse literals  | 
| 17234 | 259  | 
then error "Problem too trivial for resolution (empty clause)"  | 
260  | 
else  | 
|
| 19447 | 261  | 
     Clause {clause_id = clause_id, axiom_name = axiom_name, 
 | 
262  | 
th = th, kind = kind,  | 
|
| 
18869
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
263  | 
literals = literals, types_sorts = types_sorts};  | 
| 
17317
 
3f12de2e2e6e
Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
 
paulson 
parents: 
17312 
diff
changeset
 | 
264  | 
|
| 
 
3f12de2e2e6e
Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
 
paulson 
parents: 
17312 
diff
changeset
 | 
265  | 
|
| 
18218
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
266  | 
(*Declarations of the current theory--to allow suppressing types.*)  | 
| 
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
267  | 
val const_typargs = ref (Library.K [] : (string*typ -> typ list));  | 
| 
17317
 
3f12de2e2e6e
Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
 
paulson 
parents: 
17312 
diff
changeset
 | 
268  | 
|
| 
18218
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
269  | 
fun num_typargs(s,T) = if !keep_types then length (!const_typargs (s,T)) else 0;  | 
| 
16925
 
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
 
paulson 
parents: 
16903 
diff
changeset
 | 
270  | 
|
| 
 
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
 
paulson 
parents: 
16903 
diff
changeset
 | 
271  | 
(*Initialize the type suppression mechanism with the current theory before  | 
| 
 
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
 
paulson 
parents: 
16903 
diff
changeset
 | 
272  | 
producing any clauses!*)  | 
| 
18218
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
273  | 
fun init thy = (const_typargs := Sign.const_typargs thy);  | 
| 
16925
 
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
 
paulson 
parents: 
16903 
diff
changeset
 | 
274  | 
|
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
275  | 
|
| 
18402
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
276  | 
(*Flatten a type to a fol_type while accumulating sort constraints on the TFrees and  | 
| 
16925
 
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
 
paulson 
parents: 
16903 
diff
changeset
 | 
277  | 
TVars it contains.*)  | 
| 
18218
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
278  | 
fun type_of (Type (a, Ts)) =  | 
| 18798 | 279  | 
let val (folTyps, ts) = types_of Ts  | 
| 
18218
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
280  | 
val t = make_fixed_type_const a  | 
| 18798 | 281  | 
in (Comp(t,folTyps), ts) end  | 
282  | 
| type_of (TFree (a,s)) = (AtomF(make_fixed_type_var a), [(FOLTFree a, s)])  | 
|
283  | 
| type_of (TVar (v, s)) = (AtomV(make_schematic_type_var v), [(FOLTVar v, s)])  | 
|
| 
18218
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
284  | 
and types_of Ts =  | 
| 18798 | 285  | 
let val (folTyps,ts) = ListPair.unzip (map type_of Ts)  | 
286  | 
in (folTyps, union_all ts) end;  | 
|
| 15390 | 287  | 
|
| 
18439
 
4b517881ac7e
Added functions for fol_type; also put some functions to the signature, used by ResHolClause.
 
mengj 
parents: 
18420 
diff
changeset
 | 
288  | 
|
| 
18218
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
289  | 
fun const_types_of (c,T) = types_of (!const_typargs (c,T));  | 
| 
16925
 
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
 
paulson 
parents: 
16903 
diff
changeset
 | 
290  | 
|
| 16903 | 291  | 
(* Any variables created via the METAHYPS tactical should be treated as  | 
292  | 
universal vars, although it is represented as "Free(...)" by Isabelle *)  | 
|
293  | 
val isMeta = String.isPrefix "METAHYP1_"  | 
|
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
294  | 
|
| 18798 | 295  | 
fun pred_name_type (Const(c,T)) = (make_fixed_const c, const_types_of (c,T))  | 
| 15390 | 296  | 
| pred_name_type (Free(x,T)) =  | 
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
297  | 
      if isMeta x then raise CLAUSE("Predicate Not First Order 1", Free(x,T)) 
 | 
| 18798 | 298  | 
else (make_fixed_var x, ([],[]))  | 
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
299  | 
  | pred_name_type (v as Var _) = raise CLAUSE("Predicate Not First Order 2", v)
 | 
| 
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
300  | 
  | pred_name_type t        = raise CLAUSE("Predicate input unexpected", t);
 | 
| 15347 | 301  | 
|
| 15615 | 302  | 
|
| 
18218
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
303  | 
(* For typed equality *)  | 
| 15615 | 304  | 
(* here "arg_typ" is the type of "="'s argument's type, not the type of the equality *)  | 
305  | 
(* Find type of equality arg *)  | 
|
306  | 
fun eq_arg_type (Type("fun",[T,_])) = 
 | 
|
307  | 
let val (folT,_) = type_of T;  | 
|
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
308  | 
in folT end;  | 
| 15347 | 309  | 
|
| 18798 | 310  | 
fun fun_name_type (Const(c,T)) args = (make_fixed_const c, const_types_of (c,T))  | 
311  | 
| fun_name_type (Free(x,T)) args =  | 
|
312  | 
       if isMeta x then raise CLAUSE("Function Not First Order", Free(x,T))
 | 
|
313  | 
else (make_fixed_var x, ([],[]))  | 
|
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
314  | 
  | fun_name_type f args = raise CLAUSE("Function Not First Order 1", f);
 | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
315  | 
|
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
316  | 
(*Convert a term to a fol_term while accumulating sort constraints on the TFrees and  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
317  | 
TVars it contains.*)  | 
| 15347 | 318  | 
fun term_of (Var(ind_nm,T)) =  | 
| 18798 | 319  | 
let val (folType,ts) = type_of T  | 
320  | 
in (UVar(make_schematic_var ind_nm, folType), ts) end  | 
|
| 15347 | 321  | 
| term_of (Free(x,T)) =  | 
| 18798 | 322  | 
let val (folType, ts) = type_of T  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
323  | 
in  | 
| 18798 | 324  | 
if isMeta x then (UVar(make_schematic_var(x,0),folType), ts)  | 
325  | 
else (Fun(make_fixed_var x, [folType], []), ts)  | 
|
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
326  | 
end  | 
| 
18218
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
327  | 
| term_of app =  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
328  | 
let val (f,args) = strip_comb app  | 
| 18798 | 329  | 
val (funName,(contys,ts1)) = fun_name_type f args  | 
330  | 
val (args',ts2) = terms_of args  | 
|
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
331  | 
in  | 
| 18868 | 332  | 
(Fun(funName,contys,args'), union_all (ts1::ts2))  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
333  | 
end  | 
| 18798 | 334  | 
and terms_of ts = ListPair.unzip (map term_of ts)  | 
| 15390 | 335  | 
|
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
336  | 
(*Create a predicate value, again accumulating sort constraints.*)  | 
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
337  | 
fun pred_of (Const("op =", typ), args) =
 | 
| 
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
338  | 
let val arg_typ = eq_arg_type typ  | 
| 18798 | 339  | 
val (args',ts) = terms_of args  | 
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
340  | 
val equal_name = make_fixed_const "op ="  | 
| 
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
341  | 
in  | 
| 
18218
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
342  | 
(Predicate(equal_name,[arg_typ],args'),  | 
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
343  | 
union_all ts)  | 
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
344  | 
end  | 
| 
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
345  | 
| pred_of (pred,args) =  | 
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
346  | 
let val (pname, (predType,ts1)) = pred_name_type pred  | 
| 18798 | 347  | 
val (args',ts2) = terms_of args  | 
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
348  | 
in  | 
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
349  | 
(Predicate(pname,predType,args'), union_all (ts1::ts2))  | 
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
350  | 
end;  | 
| 15347 | 351  | 
|
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
352  | 
(*Treatment of literals, possibly negated or tagged*)  | 
| 
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
353  | 
fun predicate_of ((Const("Not",_) $ P), polarity, tag) =
 | 
| 
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
354  | 
predicate_of (P, not polarity, tag)  | 
| 
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
355  | 
  | predicate_of ((Const("HOL.tag",_) $ P), polarity, tag) =
 | 
| 
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
356  | 
predicate_of (P, polarity, true)  | 
| 
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
357  | 
| predicate_of (term,polarity,tag) =  | 
| 
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
358  | 
(pred_of (strip_comb term), polarity, tag);  | 
| 15347 | 359  | 
|
| 17888 | 360  | 
fun literals_of_term1 args (Const("Trueprop",_) $ P) = literals_of_term1 args P
 | 
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
361  | 
  | literals_of_term1 args (Const("op |",_) $ P $ Q) = 
 | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
362  | 
literals_of_term1 (literals_of_term1 args P) Q  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
363  | 
| literals_of_term1 (lits, ts) P =  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
364  | 
let val ((pred, ts'), polarity, tag) = predicate_of (P,true,false)  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
365  | 
val lits' = Literal(polarity,pred,tag) :: lits  | 
| 17234 | 366  | 
in  | 
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
367  | 
(lits', ts union ts')  | 
| 17234 | 368  | 
end;  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
369  | 
|
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
370  | 
val literals_of_term = literals_of_term1 ([],[]);  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
371  | 
|
| 
18403
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
372  | 
|
| 
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
373  | 
fun list_ord _ ([],[]) = EQUAL  | 
| 
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
374  | 
| list_ord _ ([],_) = LESS  | 
| 
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
375  | 
| list_ord _ (_,[]) = GREATER  | 
| 
19207
 
33f1b4515ce4
Tidying and tracing. Handling exn CLAUSE so that errors don't reach top-level.
 
paulson 
parents: 
19197 
diff
changeset
 | 
376  | 
| list_ord ord (x::xs, y::ys) =  | 
| 
 
33f1b4515ce4
Tidying and tracing. Handling exn CLAUSE so that errors don't reach top-level.
 
paulson 
parents: 
19197 
diff
changeset
 | 
377  | 
(case ord(x,y) of EQUAL => list_ord ord (xs,ys)  | 
| 
 
33f1b4515ce4
Tidying and tracing. Handling exn CLAUSE so that errors don't reach top-level.
 
paulson 
parents: 
19197 
diff
changeset
 | 
378  | 
| xy_ord => xy_ord);  | 
| 
 
33f1b4515ce4
Tidying and tracing. Handling exn CLAUSE so that errors don't reach top-level.
 
paulson 
parents: 
19197 
diff
changeset
 | 
379  | 
|
| 
18403
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
380  | 
fun type_ord (AtomV(_),AtomV(_)) = EQUAL  | 
| 
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
381  | 
| type_ord (AtomV(_),_) = LESS  | 
| 
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
382  | 
| type_ord (AtomF(_),AtomV(_)) = GREATER  | 
| 
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
383  | 
| type_ord (AtomF(f1),AtomF(f2)) = string_ord (f1,f2)  | 
| 
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
384  | 
| type_ord (AtomF(_),_) = LESS  | 
| 
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
385  | 
| type_ord (Comp(_,_),AtomV(_)) = GREATER  | 
| 
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
386  | 
| type_ord (Comp(_,_),AtomF(_)) = GREATER  | 
| 
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
387  | 
| type_ord (Comp(con1,args1),Comp(con2,args2)) =  | 
| 
19207
 
33f1b4515ce4
Tidying and tracing. Handling exn CLAUSE so that errors don't reach top-level.
 
paulson 
parents: 
19197 
diff
changeset
 | 
388  | 
(case string_ord(con1,con2) of EQUAL => types_ord (args1,args2)  | 
| 
 
33f1b4515ce4
Tidying and tracing. Handling exn CLAUSE so that errors don't reach top-level.
 
paulson 
parents: 
19197 
diff
changeset
 | 
389  | 
| con_ord => con_ord)  | 
| 
18403
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
390  | 
and  | 
| 18920 | 391  | 
types_ord ([],[]) = EQUAL  | 
| 
18403
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
392  | 
| types_ord (tps1,tps2) = list_ord type_ord (tps1,tps2);  | 
| 
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
393  | 
|
| 
18402
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
394  | 
|
| 18920 | 395  | 
fun term_ord (UVar _, UVar _) = EQUAL  | 
396  | 
| term_ord (UVar _, _) = LESS  | 
|
397  | 
| term_ord (Fun _, UVar _) = GREATER  | 
|
| 
18403
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
398  | 
| term_ord (Fun(f1,tps1,tms1),Fun(f2,tps2,tms2)) =  | 
| 
18420
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
399  | 
(case string_ord (f1,f2) of  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
400  | 
EQUAL =>  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
401  | 
(case terms_ord (tms1,tms2) of EQUAL => types_ord (tps1,tps2)  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
402  | 
| tms_ord => tms_ord)  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
403  | 
| fn_ord => fn_ord)  | 
| 
18403
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
404  | 
|
| 
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
405  | 
and  | 
| 18920 | 406  | 
terms_ord ([],[]) = EQUAL  | 
| 
18420
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
407  | 
| terms_ord (tms1,tms2) = list_ord term_ord (tms1,tms2);  | 
| 
18403
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
408  | 
|
| 
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
409  | 
|
| 
 
df0c0f35c897
Changed literals' ordering and the functions for sorting literals.
 
mengj 
parents: 
18402 
diff
changeset
 | 
410  | 
|
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
411  | 
fun predicate_ord (Predicate(pname1,ftyps1,ftms1),Predicate(pname2,ftyps2,ftms2)) =  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
412  | 
case string_ord (pname1,pname2) of  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
413  | 
EQUAL => (case terms_ord(ftms1,ftms2) of EQUAL => types_ord(ftyps1,ftyps2)  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
414  | 
| ftms_ord => ftms_ord)  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
415  | 
| pname_ord => pname_ord  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
416  | 
|
| 
18402
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
417  | 
|
| 
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
418  | 
fun literal_ord (Literal(false,_,_),Literal(true,_,_)) = LESS  | 
| 
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
419  | 
| literal_ord (Literal(true,_,_),Literal(false,_,_)) = GREATER  | 
| 
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
420  | 
| literal_ord (Literal(_,pred1,_),Literal(_,pred2,_)) = predicate_ord(pred1,pred2);  | 
| 
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
421  | 
|
| 
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
422  | 
fun sort_lits lits = sort literal_ord lits;  | 
| 
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
423  | 
|
| 
18420
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
424  | 
|
| 
18409
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
425  | 
(********** clause equivalence ******************)  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
426  | 
|
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
427  | 
fun check_var_pairs (x,y) [] = 0  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
428  | 
| check_var_pairs (x,y) ((u,v)::w) =  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
429  | 
if (x,y) = (u,v) then 1  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
430  | 
else  | 
| 19176 | 431  | 
if x=u orelse y=v then 2 (*conflict*)  | 
| 
18409
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
432  | 
else check_var_pairs (x,y) w;  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
433  | 
|
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
434  | 
fun type_eq (AtomV(v1),AtomV(v2)) (vars,tvars) =  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
435  | 
(case check_var_pairs (v1,v2) tvars of 0 => (true,(vars,(v1,v2)::tvars))  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
436  | 
| 1 => (true,(vars,tvars))  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
437  | 
| 2 => (false,(vars,tvars)))  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
438  | 
| type_eq (AtomV(_),_) vtvars = (false,vtvars)  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
439  | 
| type_eq (AtomF(f1),AtomF(f2)) vtvars = (f1=f2,vtvars)  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
440  | 
| type_eq (AtomF(_),_) vtvars = (false,vtvars)  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
441  | 
| type_eq (Comp(con1,args1),Comp(con2,args2)) vtvars =  | 
| 
18420
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
442  | 
let val (eq1,vtvars1) =  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
443  | 
if con1 = con2 then types_eq (args1,args2) vtvars  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
444  | 
else (false,vtvars)  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
445  | 
in  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
446  | 
(eq1,vtvars1)  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
447  | 
end  | 
| 
18409
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
448  | 
| type_eq (Comp(_,_),_) vtvars = (false,vtvars)  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
449  | 
|
| 19176 | 450  | 
and types_eq ([],[]) vtvars = (true,vtvars)  | 
| 
18420
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
451  | 
| types_eq (tp1::tps1,tp2::tps2) vtvars =  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
452  | 
let val (eq1,vtvars1) = type_eq (tp1,tp2) vtvars  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
453  | 
val (eq2,vtvars2) = if eq1 then types_eq (tps1,tps2) vtvars1  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
454  | 
else (eq1,vtvars1)  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
455  | 
in  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
456  | 
(eq2,vtvars2)  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
457  | 
end;  | 
| 
18409
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
458  | 
|
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
459  | 
|
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
460  | 
fun term_eq (UVar(v1,tp1),UVar(v2,tp2)) (vars,tvars) =  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
461  | 
(case check_var_pairs (v1,v2) vars of 0 => type_eq (tp1,tp2) (((v1,v2)::vars),tvars)  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
462  | 
| 1 => type_eq (tp1,tp2) (vars,tvars)  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
463  | 
| 2 => (false,(vars,tvars)))  | 
| 18920 | 464  | 
| term_eq (UVar _,_) vtvars = (false,vtvars)  | 
| 
18409
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
465  | 
| term_eq (Fun(f1,tps1,tms1),Fun(f2,tps2,tms2)) vtvars =  | 
| 
18420
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
466  | 
let val (eq1,vtvars1) =  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
467  | 
if f1 = f2 then terms_eq (tms1,tms2) vtvars  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
468  | 
else (false,vtvars)  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
469  | 
val (eq2,vtvars2) =  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
470  | 
if eq1 then types_eq (tps1,tps2) vtvars1  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
471  | 
else (eq1,vtvars1)  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
472  | 
in  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
473  | 
(eq2,vtvars2)  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
474  | 
end  | 
| 
18409
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
475  | 
| term_eq (Fun(_,_,_),_) vtvars = (false,vtvars)  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
476  | 
|
| 19176 | 477  | 
and terms_eq ([],[]) vtvars = (true,vtvars)  | 
| 
18420
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
478  | 
| terms_eq (tm1::tms1,tm2::tms2) vtvars =  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
479  | 
let val (eq1,vtvars1) = term_eq (tm1,tm2) vtvars  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
480  | 
val (eq2,vtvars2) = if eq1 then terms_eq (tms1,tms2) vtvars1  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
481  | 
else (eq1,vtvars1)  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
482  | 
in  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
483  | 
(eq2,vtvars2)  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
484  | 
end;  | 
| 
18409
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
485  | 
|
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
486  | 
|
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
487  | 
fun pred_eq (Predicate(pname1,tps1,tms1),Predicate(pname2,tps2,tms2)) vtvars =  | 
| 
18409
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
488  | 
let val (eq1,vtvars1) =  | 
| 19176 | 489  | 
if pname1 = pname2 then terms_eq (tms1,tms2) vtvars  | 
| 
18409
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
490  | 
else (false,vtvars)  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
491  | 
val (eq2,vtvars2) =  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
492  | 
if eq1 then types_eq (tps1,tps2) vtvars1  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
493  | 
else (eq1,vtvars1)  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
494  | 
in  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
495  | 
(eq2,vtvars2)  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
496  | 
end;  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
497  | 
|
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
498  | 
|
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
499  | 
fun lit_eq (Literal(pol1,pred1,_),Literal(pol2,pred2,_)) vtvars =  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
500  | 
if (pol1 = pol2) then pred_eq (pred1,pred2) vtvars  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
501  | 
else (false,vtvars);  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
502  | 
|
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
503  | 
fun lits_eq ([],[]) vtvars = (true,vtvars)  | 
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
504  | 
| lits_eq (l1::ls1,l2::ls2) vtvars =  | 
| 19176 | 505  | 
let val (eq1,vtvars1) = lit_eq (l1,l2) vtvars  | 
506  | 
in  | 
|
507  | 
if eq1 then lits_eq (ls1,ls2) vtvars1  | 
|
508  | 
else (false,vtvars1)  | 
|
509  | 
end  | 
|
510  | 
| lits_eq _ vtvars = (false,vtvars);  | 
|
| 
18409
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
511  | 
|
| 
18420
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
512  | 
(*Equality of two clauses up to variable renaming*)  | 
| 18798 | 513  | 
fun clause_eq (Clause{literals=lits1,...}, Clause{literals=lits2,...}) =
 | 
| 19176 | 514  | 
#1 (lits_eq (lits1,lits2) ([],[]));  | 
| 
18409
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
515  | 
|
| 
 
080094128a09
Added functions to test equivalence between clauses.
 
mengj 
parents: 
18403 
diff
changeset
 | 
516  | 
|
| 
18420
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
517  | 
(*** Hash function for clauses ***)  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
518  | 
|
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
519  | 
val xor_words = List.foldl Word.xorb 0w0;  | 
| 
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
520  | 
|
| 18920 | 521  | 
fun hashw_term (UVar _, w) = w  | 
| 18449 | 522  | 
| hashw_term (Fun(f,tps,args), w) =  | 
523  | 
List.foldl hashw_term (Polyhash.hashw_string (f,w)) args;  | 
|
| 
18420
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
524  | 
|
| 18449 | 525  | 
fun hashw_pred (Predicate(pn,_,args), w) =  | 
526  | 
List.foldl hashw_term (Polyhash.hashw_string (pn,w)) args;  | 
|
| 
18420
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
527  | 
|
| 18449 | 528  | 
fun hash1_literal (Literal(true,pred,_)) = hashw_pred (pred, 0w0)  | 
529  | 
| hash1_literal (Literal(false,pred,_)) = Word.notb (hashw_pred (pred, 0w0));  | 
|
| 
18420
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
530  | 
|
| 18798 | 531  | 
fun hash_clause (Clause{literals,...}) =
 | 
532  | 
Word.toIntX (xor_words (map hash1_literal literals));  | 
|
| 
18402
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
533  | 
|
| 
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
534  | 
|
| 18798 | 535  | 
(*Make literals for sorted type variables. FIXME: can it use map?*)  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
536  | 
fun sorts_on_typs (_, []) = ([])  | 
| 
16199
 
ee95ab217fee
no longer emits literals for type class HOL.type; also minor tidying
 
paulson 
parents: 
16039 
diff
changeset
 | 
537  | 
| sorts_on_typs (v, "HOL.type" :: s) =  | 
| 
18411
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
538  | 
sorts_on_typs (v,s) (*IGNORE sort "type"*)  | 
| 18798 | 539  | 
| sorts_on_typs ((FOLTVar indx), s::ss) =  | 
540  | 
LTVar(make_type_class s, make_schematic_type_var indx) ::  | 
|
541  | 
sorts_on_typs ((FOLTVar indx), ss)  | 
|
542  | 
| sorts_on_typs ((FOLTFree x), s::ss) =  | 
|
543  | 
LTFree(make_type_class s, make_fixed_type_var x) ::  | 
|
544  | 
sorts_on_typs ((FOLTFree x), ss);  | 
|
| 15347 | 545  | 
|
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
546  | 
|
| 18798 | 547  | 
fun pred_of_sort (LTVar (s,ty)) = (s,1)  | 
548  | 
| pred_of_sort (LTFree (s,ty)) = (s,1)  | 
|
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
549  | 
|
| 
16199
 
ee95ab217fee
no longer emits literals for type class HOL.type; also minor tidying
 
paulson 
parents: 
16039 
diff
changeset
 | 
550  | 
(*Given a list of sorted type variables, return two separate lists.  | 
| 
 
ee95ab217fee
no longer emits literals for type class HOL.type; also minor tidying
 
paulson 
parents: 
16039 
diff
changeset
 | 
551  | 
The first is for TVars, the second for TFrees.*)  | 
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
552  | 
fun add_typs_aux [] = ([],[])  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
553  | 
| add_typs_aux ((FOLTVar indx,s)::tss) =  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
554  | 
let val vs = sorts_on_typs (FOLTVar indx, s)  | 
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
555  | 
val (vss,fss) = add_typs_aux tss  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
556  | 
in  | 
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
557  | 
(vs union vss, fss)  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
558  | 
end  | 
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
559  | 
| add_typs_aux ((FOLTFree x,s)::tss) =  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
560  | 
let val fs = sorts_on_typs (FOLTFree x, s)  | 
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
561  | 
val (vss,fss) = add_typs_aux tss  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
562  | 
in  | 
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
563  | 
(vss, fs union fss)  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
564  | 
end;  | 
| 
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
565  | 
|
| 17999 | 566  | 
|
| 
18869
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
567  | 
(** make axiom and conjecture clauses. **)  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
568  | 
|
| 
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
569  | 
fun get_tvar_strs [] = []  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
570  | 
| get_tvar_strs ((FOLTVar indx,s)::tss) =  | 
| 18920 | 571  | 
(make_schematic_type_var indx) ins (get_tvar_strs tss)  | 
572  | 
| get_tvar_strs((FOLTFree x,s)::tss) = get_tvar_strs tss  | 
|
| 15347 | 573  | 
|
| 
18869
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
574  | 
(* check if a clause is first-order before making a conjecture clause*)  | 
| 
19443
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
575  | 
fun make_conjecture_clause n thm =  | 
| 
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
576  | 
let val t = prop_of thm  | 
| 
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
577  | 
val _ = check_is_fol_term t  | 
| 
18199
 
d236379ea408
-- before converting axiom and conjecture clauses into ResClause.clause format, perform "check_is_fol_term" first.
 
mengj 
parents: 
18056 
diff
changeset
 | 
578  | 
	    handle TERM("check_is_fol_term",_) => raise CLAUSE("Goal is not FOL",t)
 | 
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
579  | 
val (lits,types_sorts) = literals_of_term t  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
580  | 
in  | 
| 19447 | 581  | 
make_clause(n, "conjecture", thm, Conjecture, lits, types_sorts)  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
582  | 
end;  | 
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
583  | 
|
| 
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
584  | 
fun make_conjecture_clauses_aux _ [] = []  | 
| 17888 | 585  | 
| make_conjecture_clauses_aux n (t::ts) =  | 
586  | 
make_conjecture_clause n t :: make_conjecture_clauses_aux (n+1) ts  | 
|
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
587  | 
|
| 
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
588  | 
val make_conjecture_clauses = make_conjecture_clauses_aux 0  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
589  | 
|
| 18920 | 590  | 
(** Too general means, positive equality literal with a variable X as one operand,  | 
591  | 
when X does not occur properly in the other operand. This rules out clearly  | 
|
592  | 
inconsistent clauses such as V=a|V=b, though it by no means guarantees soundness. **)  | 
|
593  | 
||
594  | 
fun occurs a (UVar(b,_)) = a=b  | 
|
595  | 
| occurs a (Fun (_,_,ts)) = exists (occurs a) ts  | 
|
596  | 
||
597  | 
(*Is the first operand a variable that does not properly occur in the second operand?*)  | 
|
598  | 
fun too_general_terms (UVar _, UVar _) = false  | 
|
599  | 
| too_general_terms (Fun _, _) = false  | 
|
600  | 
| too_general_terms (UVar (a,_), t) = not (occurs a t);  | 
|
601  | 
||
602  | 
fun too_general_lit (Literal (true,Predicate("equal",_,[x,y]),_)) =
 | 
|
603  | 
too_general_terms (x,y) orelse too_general_terms(y,x)  | 
|
604  | 
| too_general_lit _ = false;  | 
|
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
605  | 
|
| 
18199
 
d236379ea408
-- before converting axiom and conjecture clauses into ResClause.clause format, perform "check_is_fol_term" first.
 
mengj 
parents: 
18056 
diff
changeset
 | 
606  | 
(*before converting an axiom clause to "clause" format, check if it is FOL*)  | 
| 
19443
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
607  | 
fun make_axiom_clause thm (ax_name,cls_id) =  | 
| 
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
608  | 
let val term = prop_of thm  | 
| 
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
609  | 
val (lits,types_sorts) = literals_of_term term  | 
| 15347 | 610  | 
in  | 
| 
19207
 
33f1b4515ce4
Tidying and tracing. Handling exn CLAUSE so that errors don't reach top-level.
 
paulson 
parents: 
19197 
diff
changeset
 | 
611  | 
if not (Meson.is_fol_term term) then  | 
| 
 
33f1b4515ce4
Tidying and tracing. Handling exn CLAUSE so that errors don't reach top-level.
 
paulson 
parents: 
19197 
diff
changeset
 | 
612  | 
	   (Output.debug ("Omitting " ^ ax_name ^ ": Axiom is not FOL"); 
 | 
| 
 
33f1b4515ce4
Tidying and tracing. Handling exn CLAUSE so that errors don't reach top-level.
 
paulson 
parents: 
19197 
diff
changeset
 | 
613  | 
NONE)  | 
| 
 
33f1b4515ce4
Tidying and tracing. Handling exn CLAUSE so that errors don't reach top-level.
 
paulson 
parents: 
19197 
diff
changeset
 | 
614  | 
else if forall too_general_lit lits then  | 
| 18920 | 615  | 
	   (Output.debug ("Omitting " ^ ax_name ^ ": equalities are too general"); 
 | 
616  | 
NONE)  | 
|
| 19447 | 617  | 
else SOME (make_clause(cls_id, ax_name, thm, Axiom, sort_lits lits, types_sorts))  | 
| 
19207
 
33f1b4515ce4
Tidying and tracing. Handling exn CLAUSE so that errors don't reach top-level.
 
paulson 
parents: 
19197 
diff
changeset
 | 
618  | 
end  | 
| 
 
33f1b4515ce4
Tidying and tracing. Handling exn CLAUSE so that errors don't reach top-level.
 
paulson 
parents: 
19197 
diff
changeset
 | 
619  | 
handle CLAUSE _ => NONE;  | 
| 15347 | 620  | 
|
621  | 
||
| 
19443
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
622  | 
fun make_axiom_clauses [] = []  | 
| 
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
623  | 
| make_axiom_clauses ((thm,(name,id))::thms) =  | 
| 
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
624  | 
case make_axiom_clause thm (name,id) of SOME cls => if isTaut cls then make_axiom_clauses thms else cls :: make_axiom_clauses thms  | 
| 
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
625  | 
| NONE => make_axiom_clauses thms;  | 
| 19354 | 626  | 
|
| 15347 | 627  | 
(**** Isabelle arities ****)  | 
628  | 
||
629  | 
exception ARCLAUSE of string;  | 
|
630  | 
||
631  | 
type class = string;  | 
|
632  | 
type tcons = string;  | 
|
633  | 
||
| 18868 | 634  | 
datatype arLit = TConsLit of bool * (class * tcons * string list)  | 
635  | 
| TVarLit of bool * (class * string);  | 
|
| 15347 | 636  | 
|
637  | 
datatype arityClause =  | 
|
638  | 
	 ArityClause of {clause_id: clause_id,
 | 
|
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
639  | 
axiom_name: axiom_name,  | 
| 15347 | 640  | 
kind: kind,  | 
641  | 
conclLit: arLit,  | 
|
642  | 
premLits: arLit list};  | 
|
643  | 
||
644  | 
||
| 18798 | 645  | 
fun gen_TVars 0 = []  | 
646  | 
  | gen_TVars n = ("T_" ^ Int.toString n) :: gen_TVars (n-1);
 | 
|
| 15347 | 647  | 
|
| 
18411
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
648  | 
fun pack_sort(_,[]) = []  | 
| 
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
649  | 
| pack_sort(tvar, "HOL.type"::srt) = pack_sort(tvar, srt) (*IGNORE sort "type"*)  | 
| 
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
650  | 
| pack_sort(tvar, cls::srt) = (make_type_class cls, tvar) :: pack_sort(tvar, srt);  | 
| 15347 | 651  | 
|
| 18868 | 652  | 
fun make_TVarLit (b, (cls,str)) = TVarLit(b, (cls,str));  | 
653  | 
fun make_TConsLit (b, (cls,tcons,tvars)) =  | 
|
654  | 
TConsLit(b, (make_type_class cls, make_fixed_type_const tcons, tvars));  | 
|
| 15347 | 655  | 
|
| 
18411
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
656  | 
(*Arity of type constructor tcon :: (arg1,...,argN)res*)  | 
| 
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
657  | 
fun make_axiom_arity_clause (tcons, n, (res,args)) =  | 
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
658  | 
let val nargs = length args  | 
| 18798 | 659  | 
val tvars = gen_TVars nargs  | 
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
660  | 
val tvars_srts = ListPair.zip (tvars,args)  | 
| 
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
661  | 
val tvars_srts' = union_all(map pack_sort tvars_srts)  | 
| 
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
662  | 
val false_tvars_srts' = map (pair false) tvars_srts'  | 
| 
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
663  | 
in  | 
| 
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
664  | 
      ArityClause {clause_id = n, kind = Axiom, 
 | 
| 
18411
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
665  | 
axiom_name = lookup_type_const tcons,  | 
| 
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
666  | 
conclLit = make_TConsLit(true, (res,tcons,tvars)),  | 
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
667  | 
premLits = map make_TVarLit false_tvars_srts'}  | 
| 
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
668  | 
end;  | 
| 15347 | 669  | 
|
670  | 
||
671  | 
(**** Isabelle class relations ****)  | 
|
672  | 
||
673  | 
datatype classrelClause =  | 
|
| 18868 | 674  | 
	 ClassrelClause of {axiom_name: axiom_name,
 | 
| 15347 | 675  | 
subclass: class,  | 
| 
18411
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
676  | 
superclass: class};  | 
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
677  | 
|
| 
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
678  | 
fun make_axiom_classrelClause n subclass superclass =  | 
| 18868 | 679  | 
  ClassrelClause {axiom_name = clrelclause_prefix ^ ascii_of subclass ^ 
 | 
680  | 
"_" ^ Int.toString n,  | 
|
681  | 
subclass = make_type_class subclass,  | 
|
682  | 
superclass = make_type_class superclass};  | 
|
| 15347 | 683  | 
|
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
684  | 
fun classrelClauses_of_aux n sub [] = []  | 
| 
18411
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
685  | 
  | classrelClauses_of_aux n sub ("HOL.type"::sups) = (*Should be ignored*)
 | 
| 
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
686  | 
classrelClauses_of_aux n sub sups  | 
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
687  | 
| classrelClauses_of_aux n sub (sup::sups) =  | 
| 18868 | 688  | 
make_axiom_classrelClause n sub sup :: classrelClauses_of_aux (n+1) sub sups;  | 
| 15347 | 689  | 
|
| 
18411
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
690  | 
fun classrelClauses_of (sub,sups) = classrelClauses_of_aux 0 sub sups;  | 
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
691  | 
|
| 18868 | 692  | 
val classrel_of = #2 o #classes o Type.rep_tsig o Sign.tsig_of;  | 
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
693  | 
|
| 18868 | 694  | 
fun classrel_clauses_classrel (C: Sorts.classes) =  | 
695  | 
map classrelClauses_of (Graph.dest C);  | 
|
696  | 
||
697  | 
val classrel_clauses_thy = List.concat o classrel_clauses_classrel o classrel_of;  | 
|
698  | 
||
699  | 
||
700  | 
(** Isabelle arities **)  | 
|
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
701  | 
|
| 
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
702  | 
fun arity_clause _ (tcons, []) = []  | 
| 19155 | 703  | 
  | arity_clause n (tcons, ("HOL.type",_)::ars) =  (*ignore*)
 | 
| 
18411
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
704  | 
arity_clause n (tcons,ars)  | 
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
705  | 
| arity_clause n (tcons, ar::ars) =  | 
| 
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
706  | 
make_axiom_arity_clause (tcons,n,ar) ::  | 
| 
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
707  | 
arity_clause (n+1) (tcons,ars);  | 
| 
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
708  | 
|
| 
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
709  | 
fun multi_arity_clause [] = []  | 
| 19155 | 710  | 
| multi_arity_clause ((tcons,ars) :: tc_arlists) =  | 
711  | 
(*Reversal ensures that older entries always get the same axiom name*)  | 
|
712  | 
arity_clause 0 (tcons, rev ars) @  | 
|
713  | 
multi_arity_clause tc_arlists  | 
|
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
714  | 
|
| 
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
715  | 
fun arity_clause_thy thy =  | 
| 19521 | 716  | 
let val arities =  | 
717  | 
Symtab.dest (Sign.arities_of thy) |> map (apsnd (map (fn (c, (_, Ss)) => (c, Ss))));  | 
|
718  | 
in multi_arity_clause (rev arities) end;  | 
|
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
719  | 
|
| 
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
720  | 
|
| 18868 | 721  | 
(**** Find occurrences of predicates in clauses ****)  | 
722  | 
||
723  | 
(*FIXME: multiple-arity checking doesn't work, as update_new is the wrong  | 
|
724  | 
function (it flags repeated declarations of a function, even with the same arity)*)  | 
|
725  | 
||
726  | 
fun update_many (tab, keypairs) = foldl (uncurry Symtab.update) tab keypairs;  | 
|
727  | 
||
728  | 
fun add_predicate_preds (Predicate(pname,tys,tms), preds) =  | 
|
729  | 
if pname = "equal" then preds (*equality is built-in and requires no declaration*)  | 
|
730  | 
else Symtab.update (pname, length tys + length tms) preds  | 
|
731  | 
||
732  | 
fun add_literal_preds (Literal(_,pred,_), preds) = add_predicate_preds (pred,preds)  | 
|
733  | 
||
734  | 
fun add_type_sort_preds ((FOLTVar indx,s), preds) =  | 
|
735  | 
update_many (preds, map pred_of_sort (sorts_on_typs (FOLTVar indx, s)))  | 
|
736  | 
| add_type_sort_preds ((FOLTFree x,s), preds) =  | 
|
737  | 
update_many (preds, map pred_of_sort (sorts_on_typs (FOLTFree x, s)));  | 
|
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
738  | 
|
| 18868 | 739  | 
fun add_clause_preds (Clause {literals, types_sorts, ...}, preds) =
 | 
740  | 
foldl add_literal_preds (foldl add_type_sort_preds preds types_sorts) literals  | 
|
741  | 
  handle Symtab.DUP a => raise ERROR ("predicate " ^ a ^ " has multiple arities")
 | 
|
742  | 
||
743  | 
fun add_classrelClause_preds (ClassrelClause {subclass,superclass,...}, preds) =
 | 
|
744  | 
Symtab.update (subclass,1) (Symtab.update (superclass,1) preds);  | 
|
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
745  | 
|
| 18868 | 746  | 
fun add_arityClause_preds (ArityClause {conclLit,...}, preds) =
 | 
747  | 
let val TConsLit(_, (tclass, _, _)) = conclLit  | 
|
748  | 
in Symtab.update (tclass,1) preds end;  | 
|
749  | 
||
750  | 
fun preds_of_clauses clauses clsrel_clauses arity_clauses =  | 
|
751  | 
Symtab.dest  | 
|
752  | 
(foldl add_classrelClause_preds  | 
|
753  | 
(foldl add_arityClause_preds  | 
|
754  | 
(foldl add_clause_preds Symtab.empty clauses)  | 
|
755  | 
arity_clauses)  | 
|
756  | 
clsrel_clauses)  | 
|
| 18798 | 757  | 
|
| 18868 | 758  | 
(*** Find occurrences of functions in clauses ***)  | 
759  | 
||
760  | 
fun add_foltype_funcs (AtomV _, funcs) = funcs  | 
|
761  | 
| add_foltype_funcs (AtomF a, funcs) = Symtab.update (a,0) funcs  | 
|
762  | 
| add_foltype_funcs (Comp(a,tys), funcs) =  | 
|
763  | 
foldl add_foltype_funcs (Symtab.update (a, length tys) funcs) tys;  | 
|
764  | 
||
765  | 
fun add_folterm_funcs (UVar _, funcs) = funcs  | 
|
766  | 
| add_folterm_funcs (Fun(a,tys,[]), funcs) = Symtab.update (a,0) funcs  | 
|
767  | 
(*A constant is a special case: it has no type argument even if overloaded*)  | 
|
768  | 
| add_folterm_funcs (Fun(a,tys,tms), funcs) =  | 
|
769  | 
foldl add_foltype_funcs  | 
|
770  | 
(foldl add_folterm_funcs (Symtab.update (a, length tys + length tms) funcs)  | 
|
771  | 
tms)  | 
|
772  | 
tys  | 
|
| 18798 | 773  | 
|
| 18868 | 774  | 
fun add_predicate_funcs (Predicate(_,tys,tms), funcs) =  | 
775  | 
foldl add_foltype_funcs (foldl add_folterm_funcs funcs tms) tys;  | 
|
776  | 
||
777  | 
fun add_literal_funcs (Literal(_,pred,_), funcs) = add_predicate_funcs (pred,funcs)  | 
|
778  | 
||
779  | 
fun add_arityClause_funcs (ArityClause {conclLit,...}, funcs) =
 | 
|
780  | 
let val TConsLit(_, (_, tcons, tvars)) = conclLit  | 
|
781  | 
in Symtab.update (tcons, length tvars) funcs end;  | 
|
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
782  | 
|
| 18868 | 783  | 
fun add_clause_funcs (Clause {literals, ...}, funcs) =
 | 
784  | 
foldl add_literal_funcs funcs literals  | 
|
785  | 
  handle Symtab.DUP a => raise ERROR ("function " ^ a ^ " has multiple arities")
 | 
|
786  | 
||
787  | 
fun funcs_of_clauses clauses arity_clauses =  | 
|
788  | 
Symtab.dest (foldl add_arityClause_funcs  | 
|
789  | 
(foldl add_clause_funcs Symtab.empty clauses)  | 
|
790  | 
arity_clauses)  | 
|
791  | 
||
792  | 
||
793  | 
(**** String-oriented operations ****)  | 
|
| 15347 | 794  | 
|
795  | 
fun wrap_eq_type typ t = eq_typ_wrapper ^"(" ^ t ^ "," ^ typ ^ ")";
 | 
|
796  | 
||
| 
18218
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
797  | 
(*Only need to wrap equality's arguments with "typeinfo" if the output clauses are typed  | 
| 
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
798  | 
and if we specifically ask for types to be included. *)  | 
| 15347 | 799  | 
fun string_of_equality (typ,terms) =  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
800  | 
let val [tstr1,tstr2] = map string_of_term terms  | 
| 
18402
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
801  | 
val typ' = string_of_fol_type typ  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
802  | 
in  | 
| 
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
803  | 
if !keep_types andalso !special_equal  | 
| 
18402
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
804  | 
	  then "equal(" ^ (wrap_eq_type typ' tstr1) ^ "," ^ 
 | 
| 
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
805  | 
(wrap_eq_type typ' tstr2) ^ ")"  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
806  | 
	  else "equal(" ^ tstr1 ^ "," ^ tstr2 ^ ")"
 | 
| 
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
807  | 
end  | 
| 
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
808  | 
and string_of_term (UVar(x,_)) = x  | 
| 
18218
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
809  | 
  | string_of_term (Fun("equal",[typ],terms)) = string_of_equality(typ,terms)
 | 
| 
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
810  | 
| string_of_term (Fun (name,typs,[])) = name (*Overloaded consts like 0 don't get types!*)  | 
| 
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
811  | 
| string_of_term (Fun (name,typs,terms)) =  | 
| 
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
812  | 
let val terms_as_strings = map string_of_term terms  | 
| 
18402
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
813  | 
val typs' = if !keep_types then map string_of_fol_type typs else []  | 
| 
18420
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
814  | 
in name ^ (paren_pack (terms_as_strings @ typs')) end;  | 
| 15347 | 815  | 
|
816  | 
(* before output the string of the predicate, check if the predicate corresponds to an equality or not. *)  | 
|
| 
18218
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
817  | 
fun string_of_predicate (Predicate("equal",[typ],terms)) = string_of_equality(typ,terms)
 | 
| 
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
818  | 
| string_of_predicate (Predicate(name,typs,terms)) =  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
819  | 
let val terms_as_strings = map string_of_term terms  | 
| 
18402
 
aaba095cf62b
1. changed fol_type, it's not a string type anymore.
 
mengj 
parents: 
18390 
diff
changeset
 | 
820  | 
val typs' = if !keep_types then map string_of_fol_type typs else []  | 
| 
18420
 
9470061ab283
hashing to eliminate the output of duplicate clauses
 
paulson 
parents: 
18411 
diff
changeset
 | 
821  | 
in name ^ (paren_pack (terms_as_strings @ typs')) end;  | 
| 
17317
 
3f12de2e2e6e
Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
 
paulson 
parents: 
17312 
diff
changeset
 | 
822  | 
|
| 
 
3f12de2e2e6e
Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
 
paulson 
parents: 
17312 
diff
changeset
 | 
823  | 
fun string_of_clausename (cls_id,ax_name) =  | 
| 
17525
 
ae5bb6001afb
tidying, and support for axclass/classrel clauses
 
paulson 
parents: 
17422 
diff
changeset
 | 
824  | 
clause_prefix ^ ascii_of ax_name ^ "_" ^ Int.toString cls_id;  | 
| 
17317
 
3f12de2e2e6e
Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
 
paulson 
parents: 
17312 
diff
changeset
 | 
825  | 
|
| 
 
3f12de2e2e6e
Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
 
paulson 
parents: 
17312 
diff
changeset
 | 
826  | 
fun string_of_type_clsname (cls_id,ax_name,idx) =  | 
| 
17525
 
ae5bb6001afb
tidying, and support for axclass/classrel clauses
 
paulson 
parents: 
17422 
diff
changeset
 | 
827  | 
string_of_clausename (cls_id,ax_name) ^ "_tcs" ^ (Int.toString idx);  | 
| 18863 | 828  | 
|
829  | 
(*Write a list of strings to a file*)  | 
|
830  | 
fun writeln_strs os = List.app (fn s => TextIO.output (os,s));  | 
|
831  | 
||
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
832  | 
|
| 18868 | 833  | 
(**** Producing DFG files ****)  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
834  | 
|
| 18863 | 835  | 
(*Attach sign in DFG syntax: false means negate.*)  | 
836  | 
fun dfg_sign true s = s  | 
|
837  | 
  | dfg_sign false s = "not(" ^ s ^ ")"  
 | 
|
838  | 
||
839  | 
fun dfg_literal (Literal(pol,pred,tag)) = dfg_sign pol (string_of_predicate pred)  | 
|
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
840  | 
|
| 18798 | 841  | 
fun dfg_of_typeLit (LTVar (s,ty)) = "not(" ^ s ^ "(" ^ ty ^ "))"
 | 
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
842  | 
  | dfg_of_typeLit (LTFree (s,ty)) = s ^ "(" ^ ty ^ ")";
 | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
843  | 
|
| 18868 | 844  | 
(*Enclose the clause body by quantifiers, if necessary*)  | 
845  | 
fun dfg_forall [] body = body  | 
|
846  | 
| dfg_forall vars body = "forall([" ^ commas vars ^ "],\n" ^ body ^ ")"  | 
|
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
847  | 
|
| 18868 | 848  | 
fun gen_dfg_cls (cls_id, ax_name, knd, lits, vars) =  | 
849  | 
    "clause( %(" ^ knd ^ ")\n" ^ 
 | 
|
850  | 
    dfg_forall vars ("or(" ^ lits ^ ")") ^ ",\n" ^ 
 | 
|
| 18863 | 851  | 
string_of_clausename (cls_id,ax_name) ^ ").\n\n";  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
852  | 
|
| 
18869
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
853  | 
fun dfg_clause_aux (Clause{literals, types_sorts, ...}) = 
 | 
| 18868 | 854  | 
let val lits = map dfg_literal literals  | 
| 
18869
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
855  | 
val (tvar_lits,tfree_lits) = add_typs_aux types_sorts  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
856  | 
val tvar_lits_strs =  | 
| 
18869
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
857  | 
if !keep_types then map dfg_of_typeLit tvar_lits else []  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
858  | 
val tfree_lits =  | 
| 
18869
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
859  | 
if !keep_types then map dfg_of_typeLit tfree_lits else []  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
860  | 
in  | 
| 17234 | 861  | 
(tvar_lits_strs @ lits, tfree_lits)  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
862  | 
end;  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
863  | 
|
| 
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
864  | 
fun dfg_folterms (Literal(pol,pred,tag)) =  | 
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
865  | 
let val Predicate (_, _, folterms) = pred  | 
| 
18218
 
9a7ffce389c3
new treatment of polymorphic types, using Sign.const_typargs
 
paulson 
parents: 
18199 
diff
changeset
 | 
866  | 
in folterms end  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
867  | 
|
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
868  | 
fun get_uvars (UVar(a,typ)) = [a]  | 
| 18868 | 869  | 
| get_uvars (Fun (_,typ,tlist)) = union_all(map get_uvars tlist)  | 
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17375 
diff
changeset
 | 
870  | 
|
| 18868 | 871  | 
fun dfg_vars (Clause {literals,...}) =
 | 
| 18920 | 872  | 
union_all (map get_uvars (List.concat (map dfg_folterms literals)))  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
873  | 
|
| 18798 | 874  | 
fun clause2dfg (cls as Clause{axiom_name,clause_id,kind,types_sorts,...}) =
 | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
875  | 
let val (lits,tfree_lits) = dfg_clause_aux cls  | 
| 18798 | 876  | 
(*"lits" includes the typing assumptions (TVars)*)  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
877  | 
val vars = dfg_vars cls  | 
| 18798 | 878  | 
val tvars = get_tvar_strs types_sorts  | 
879  | 
val knd = name_of_kind kind  | 
|
| 17234 | 880  | 
val lits_str = commas lits  | 
| 18868 | 881  | 
val cls_str = gen_dfg_cls(clause_id, axiom_name, knd, lits_str, tvars@vars)  | 
| 18798 | 882  | 
in (cls_str, tfree_lits) end;  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
883  | 
|
| 18798 | 884  | 
fun string_of_arity (name, num) =  "(" ^ name ^ "," ^ Int.toString num ^ ")"
 | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
885  | 
|
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
886  | 
fun string_of_preds [] = ""  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
887  | 
| string_of_preds preds = "predicates[" ^ commas(map string_of_arity preds) ^ "].\n";  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
888  | 
|
| 
18856
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
889  | 
fun string_of_funcs [] = ""  | 
| 
 
4669dec681f4
tidy-up of res_clause.ML, removing the "predicates" field
 
paulson 
parents: 
18798 
diff
changeset
 | 
890  | 
| string_of_funcs funcs = "functions[" ^ commas(map string_of_arity funcs) ^ "].\n" ;  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
891  | 
|
| 17234 | 892  | 
fun string_of_symbols predstr funcstr =  | 
893  | 
"list_of_symbols.\n" ^ predstr ^ funcstr ^ "end_of_list.\n\n";  | 
|
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
894  | 
|
| 18798 | 895  | 
fun string_of_start name = "begin_problem(" ^ name ^ ").\n\n";
 | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
896  | 
|
| 18863 | 897  | 
fun string_of_descrip name =  | 
| 18868 | 898  | 
  "list_of_descriptions.\nname({*" ^ name ^ 
 | 
899  | 
  "*}).\nauthor({*Isabelle*}).\nstatus(unknown).\ndescription({*auto-generated*}).\nend_of_list.\n\n"
 | 
|
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
900  | 
|
| 18863 | 901  | 
fun dfg_tfree_clause tfree_lit =  | 
902  | 
"clause( %(conjecture)\n" ^ "or( " ^ tfree_lit ^ "),\n" ^ "tfree_tcs" ^ ").\n\n"  | 
|
903  | 
||
| 
17845
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
904  | 
fun string_of_arClauseID (ArityClause {clause_id,axiom_name,...}) =
 | 
| 
 
1438291d57f0
deletion of Tools/res_types_sorts; removal of absolute numbering of clauses
 
paulson 
parents: 
17775 
diff
changeset
 | 
905  | 
arclause_prefix ^ ascii_of axiom_name ^ "_" ^ Int.toString clause_id;  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
906  | 
|
| 18863 | 907  | 
fun dfg_of_arLit (TConsLit(pol,(c,t,args))) =  | 
908  | 
      dfg_sign pol (c ^ "(" ^ t ^ paren_pack args ^ ")")
 | 
|
909  | 
| dfg_of_arLit (TVarLit(pol,(c,str))) =  | 
|
910  | 
      dfg_sign pol (c ^ "(" ^ str ^ ")")
 | 
|
| 
17525
 
ae5bb6001afb
tidying, and support for axclass/classrel clauses
 
paulson 
parents: 
17422 
diff
changeset
 | 
911  | 
|
| 18863 | 912  | 
fun dfg_classrelLits sub sup =  | 
913  | 
let val tvar = "(T)"  | 
|
914  | 
in  | 
|
915  | 
	"not(" ^ sub ^ tvar ^ "), " ^ sup ^ tvar
 | 
|
916  | 
end;  | 
|
| 
17525
 
ae5bb6001afb
tidying, and support for axclass/classrel clauses
 
paulson 
parents: 
17422 
diff
changeset
 | 
917  | 
|
| 18868 | 918  | 
fun dfg_classrelClause (ClassrelClause {axiom_name,subclass,superclass,...}) =
 | 
919  | 
"clause(forall([T],\nor( " ^ dfg_classrelLits subclass superclass ^ ")),\n" ^  | 
|
920  | 
axiom_name ^ ").\n\n";  | 
|
921  | 
||
922  | 
fun dfg_arity_clause (arcls as ArityClause{kind,conclLit,premLits,...}) = 
 | 
|
923  | 
let val arcls_id = string_of_arClauseID arcls  | 
|
924  | 
val knd = name_of_kind kind  | 
|
925  | 
val TConsLit(_, (_,_,tvars)) = conclLit  | 
|
926  | 
val lits = map dfg_of_arLit (conclLit :: premLits)  | 
|
| 18863 | 927  | 
in  | 
| 18868 | 928  | 
      "clause( %(" ^ knd ^ ")\n" ^ 
 | 
929  | 
      dfg_forall tvars ("or( " ^ commas lits ^ ")") ^ ",\n" ^
 | 
|
930  | 
arcls_id ^ ").\n\n"  | 
|
| 18863 | 931  | 
end;  | 
932  | 
||
933  | 
(* write out a subgoal in DFG format to the file "xxxx_N"*)  | 
|
934  | 
fun dfg_write_file ths filename (axclauses,classrel_clauses,arity_clauses) =  | 
|
935  | 
let  | 
|
| 
19207
 
33f1b4515ce4
Tidying and tracing. Handling exn CLAUSE so that errors don't reach top-level.
 
paulson 
parents: 
19197 
diff
changeset
 | 
936  | 
    val _ = Output.debug ("Preparing to write the DFG file " ^ filename)
 | 
| 
19443
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
937  | 
val conjectures = make_conjecture_clauses ths  | 
| 18868 | 938  | 
val (dfg_clss, tfree_litss) = ListPair.unzip (map clause2dfg conjectures)  | 
| 18863 | 939  | 
val clss = conjectures @ axclauses  | 
| 18868 | 940  | 
val funcs = funcs_of_clauses clss arity_clauses  | 
941  | 
and preds = preds_of_clauses clss classrel_clauses arity_clauses  | 
|
| 18863 | 942  | 
and probname = Path.pack (Path.base (Path.unpack filename))  | 
| 18868 | 943  | 
val (axstrs, _) = ListPair.unzip (map clause2dfg axclauses)  | 
| 18863 | 944  | 
val tfree_clss = map dfg_tfree_clause (union_all tfree_litss)  | 
| 19155 | 945  | 
val out = TextIO.openOut filename  | 
| 18863 | 946  | 
in  | 
| 18868 | 947  | 
TextIO.output (out, string_of_start probname);  | 
948  | 
TextIO.output (out, string_of_descrip probname);  | 
|
949  | 
TextIO.output (out, string_of_symbols (string_of_funcs funcs) (string_of_preds preds));  | 
|
950  | 
TextIO.output (out, "list_of_clauses(axioms,cnf).\n");  | 
|
951  | 
writeln_strs out axstrs;  | 
|
952  | 
List.app (curry TextIO.output out o dfg_classrelClause) classrel_clauses;  | 
|
953  | 
List.app (curry TextIO.output out o dfg_arity_clause) arity_clauses;  | 
|
954  | 
TextIO.output (out, "end_of_list.\n\nlist_of_clauses(conjectures,cnf).\n");  | 
|
955  | 
writeln_strs out tfree_clss;  | 
|
956  | 
writeln_strs out dfg_clss;  | 
|
957  | 
TextIO.output (out, "end_of_list.\n\nend_problem.\n");  | 
|
958  | 
TextIO.closeOut out  | 
|
| 
17525
 
ae5bb6001afb
tidying, and support for axclass/classrel clauses
 
paulson 
parents: 
17422 
diff
changeset
 | 
959  | 
end;  | 
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
960  | 
|
| 
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
961  | 
|
| 
18869
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
962  | 
(**** Produce TPTP files ****)  | 
| 18868 | 963  | 
|
964  | 
(*Attach sign in TPTP syntax: false means negate.*)  | 
|
965  | 
fun tptp_sign true s = "++" ^ s  | 
|
966  | 
| tptp_sign false s = "--" ^ s  | 
|
967  | 
||
968  | 
fun tptp_literal (Literal(pol,pred,tag)) = (*FIXME REMOVE TAGGING*)  | 
|
| 15347 | 969  | 
let val pred_string = string_of_predicate pred  | 
| 
17230
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
970  | 
val tagged_pol =  | 
| 
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
971  | 
if (tag andalso !tagged) then (if pol then "+++" else "---")  | 
| 
 
77e93bf303a5
fixed arities and restored changes that had gone missing
 
paulson 
parents: 
17150 
diff
changeset
 | 
972  | 
else (if pol then "++" else "--")  | 
| 15347 | 973  | 
in  | 
974  | 
tagged_pol ^ pred_string  | 
|
975  | 
end;  | 
|
976  | 
||
| 18798 | 977  | 
fun tptp_of_typeLit (LTVar (s,ty)) = "--" ^ s ^ "(" ^ ty ^ ")"
 | 
978  | 
  | tptp_of_typeLit (LTFree (s,ty)) = "++" ^ s ^ "(" ^ ty ^ ")";
 | 
|
| 15347 | 979  | 
|
980  | 
fun gen_tptp_cls (cls_id,ax_name,knd,lits) =  | 
|
| 
17317
 
3f12de2e2e6e
Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
 
paulson 
parents: 
17312 
diff
changeset
 | 
981  | 
    "input_clause(" ^ string_of_clausename (cls_id,ax_name) ^ "," ^ 
 | 
| 18863 | 982  | 
knd ^ "," ^ lits ^ ").\n";  | 
| 15347 | 983  | 
|
| 
17317
 
3f12de2e2e6e
Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
 
paulson 
parents: 
17312 
diff
changeset
 | 
984  | 
fun gen_tptp_type_cls (cls_id,ax_name,knd,tfree_lit,idx) =  | 
| 
 
3f12de2e2e6e
Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
 
paulson 
parents: 
17312 
diff
changeset
 | 
985  | 
    "input_clause(" ^ string_of_type_clsname (cls_id,ax_name,idx) ^ "," ^ 
 | 
| 18863 | 986  | 
knd ^ ",[" ^ tfree_lit ^ "]).\n";  | 
| 15347 | 987  | 
|
| 
18869
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
988  | 
fun tptp_type_lits (Clause {literals, types_sorts, ...}) = 
 | 
| 
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
989  | 
let val lits = map tptp_literal literals  | 
| 
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
990  | 
val (tvar_lits,tfree_lits) = add_typs_aux types_sorts  | 
| 
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
991  | 
val tvar_lits_strs =  | 
| 
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
992  | 
if !keep_types then map tptp_of_typeLit tvar_lits else []  | 
| 
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
993  | 
val tfree_lits =  | 
| 
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
994  | 
if !keep_types then map tptp_of_typeLit tfree_lits else []  | 
| 15347 | 995  | 
in  | 
| 
17305
 
6cef3aedd661
axioms  now included in tptp files, no /bin/cat and various tidying
 
paulson 
parents: 
17261 
diff
changeset
 | 
996  | 
(tvar_lits_strs @ lits, tfree_lits)  | 
| 15347 | 997  | 
end;  | 
998  | 
||
| 
18869
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
999  | 
fun clause2tptp (cls as Clause {clause_id, axiom_name, kind, ...}) =
 | 
| 17422 | 1000  | 
let val (lits,tfree_lits) = tptp_type_lits cls  | 
| 
17305
 
6cef3aedd661
axioms  now included in tptp files, no /bin/cat and various tidying
 
paulson 
parents: 
17261 
diff
changeset
 | 
1001  | 
(*"lits" includes the typing assumptions (TVars)*)  | 
| 
18869
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
1002  | 
val knd = name_of_kind kind  | 
| 
 
00741f7280f7
removal of ResClause.num_of_clauses and other simplifications
 
paulson 
parents: 
18868 
diff
changeset
 | 
1003  | 
val cls_str = gen_tptp_cls(clause_id, axiom_name, knd, bracket_pack lits)  | 
| 15608 | 1004  | 
in  | 
1005  | 
(cls_str,tfree_lits)  | 
|
1006  | 
end;  | 
|
1007  | 
||
| 18863 | 1008  | 
fun tptp_tfree_clause tfree_lit =  | 
1009  | 
    "input_clause(" ^ "tfree_tcs," ^ "conjecture" ^ ",[" ^ tfree_lit ^ "]).\n";
 | 
|
| 15608 | 1010  | 
|
| 
17525
 
ae5bb6001afb
tidying, and support for axclass/classrel clauses
 
paulson 
parents: 
17422 
diff
changeset
 | 
1011  | 
fun tptp_of_arLit (TConsLit(b,(c,t,args))) =  | 
| 18868 | 1012  | 
      tptp_sign b (c ^ "(" ^ t ^ paren_pack args ^ ")")
 | 
| 
17525
 
ae5bb6001afb
tidying, and support for axclass/classrel clauses
 
paulson 
parents: 
17422 
diff
changeset
 | 
1013  | 
| tptp_of_arLit (TVarLit(b,(c,str))) =  | 
| 18868 | 1014  | 
      tptp_sign b (c ^ "(" ^ str ^ ")")
 | 
| 15347 | 1015  | 
|
| 18868 | 1016  | 
fun tptp_arity_clause (arcls as ArityClause{kind,conclLit,premLits,...}) = 
 | 
1017  | 
let val arcls_id = string_of_arClauseID arcls  | 
|
1018  | 
val knd = name_of_kind kind  | 
|
1019  | 
val lits = map tptp_of_arLit (conclLit :: premLits)  | 
|
1020  | 
in  | 
|
1021  | 
    "input_clause(" ^ arcls_id ^ "," ^ knd ^ "," ^ bracket_pack lits ^ ").\n"
 | 
|
1022  | 
end;  | 
|
| 15347 | 1023  | 
|
1024  | 
fun tptp_classrelLits sub sup =  | 
|
1025  | 
let val tvar = "(T)"  | 
|
1026  | 
in  | 
|
| 
18411
 
2d3165a0fb40
No axiom, arity or classrel clauses contain HOL.type. Negative occurrences are
 
paulson 
parents: 
18409 
diff
changeset
 | 
1027  | 
"[--" ^ sub ^ tvar ^ ",++" ^ sup ^ tvar ^ "]"  | 
| 15347 | 1028  | 
end;  | 
1029  | 
||
| 18868 | 1030  | 
fun tptp_classrelClause (ClassrelClause {axiom_name,subclass,superclass,...}) =
 | 
1031  | 
  "input_clause(" ^ axiom_name ^ ",axiom," ^ tptp_classrelLits subclass superclass ^ ").\n" 
 | 
|
| 
17150
 
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
 
quigley 
parents: 
16976 
diff
changeset
 | 
1032  | 
|
| 18863 | 1033  | 
(* write out a subgoal as tptp clauses to the file "xxxx_N"*)  | 
| 
19443
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
1034  | 
fun tptp_write_file thms filename (axclauses,classrel_clauses,arity_clauses) =  | 
| 18863 | 1035  | 
let  | 
| 
19207
 
33f1b4515ce4
Tidying and tracing. Handling exn CLAUSE so that errors don't reach top-level.
 
paulson 
parents: 
19197 
diff
changeset
 | 
1036  | 
    val _ = Output.debug ("Preparing to write the TPTP file " ^ filename)
 | 
| 
19443
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
1037  | 
val clss = make_conjecture_clauses thms  | 
| 
 
e32a4703d834
Take conjectures and axioms as thms when convert them to ResClause.clause format.
 
mengj 
parents: 
19354 
diff
changeset
 | 
1038  | 
val axclauses' = make_axiom_clauses axclauses  | 
| 18863 | 1039  | 
val (tptp_clss,tfree_litss) = ListPair.unzip (map clause2tptp clss)  | 
1040  | 
val tfree_clss = map tptp_tfree_clause (foldl (op union_string) [] tfree_litss)  | 
|
1041  | 
val out = TextIO.openOut filename  | 
|
1042  | 
in  | 
|
| 
19197
 
92404b5c20ad
tptp_write_file now takes goals and axioms as Term.term and writes them to a file.
 
mengj 
parents: 
19176 
diff
changeset
 | 
1043  | 
List.app (curry TextIO.output out o #1 o clause2tptp) axclauses';  | 
| 18863 | 1044  | 
writeln_strs out tfree_clss;  | 
1045  | 
writeln_strs out tptp_clss;  | 
|
| 18868 | 1046  | 
List.app (curry TextIO.output out o tptp_classrelClause) classrel_clauses;  | 
1047  | 
List.app (curry TextIO.output out o tptp_arity_clause) arity_clauses;  | 
|
| 18863 | 1048  | 
TextIO.closeOut out  | 
1049  | 
end;  | 
|
1050  | 
||
| 19354 | 1051  | 
|
1052  | 
||
1053  | 
||
1054  | 
||
| 15347 | 1055  | 
end;  |