| author | wenzelm | 
| Fri, 07 Jan 2011 17:07:00 +0100 | |
| changeset 41447 | 537b290bbe38 | 
| parent 38795 | 848be46708dc | 
| child 41471 | 54a58904a598 | 
| permissions | -rw-r--r-- | 
| 14452 | 1  | 
(* Title: HOL/Tools/prop_logic.ML  | 
2  | 
Author: Tjark Weber  | 
|
| 31220 | 3  | 
Copyright 2004-2009  | 
| 14452 | 4  | 
|
5  | 
Formulas of propositional logic.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature PROP_LOGIC =  | 
|
9  | 
sig  | 
|
| 41447 | 10  | 
datatype prop_formula =  | 
11  | 
True  | 
|
12  | 
| False  | 
|
13  | 
| BoolVar of int (* NOTE: only use indices >= 1 *)  | 
|
14  | 
| Not of prop_formula  | 
|
15  | 
| Or of prop_formula * prop_formula  | 
|
16  | 
| And of prop_formula * prop_formula  | 
|
| 14452 | 17  | 
|
| 41447 | 18  | 
val SNot: prop_formula -> prop_formula  | 
19  | 
val SOr: prop_formula * prop_formula -> prop_formula  | 
|
20  | 
val SAnd: prop_formula * prop_formula -> prop_formula  | 
|
21  | 
val simplify: prop_formula -> prop_formula (* eliminates True/False and double-negation *)  | 
|
| 14452 | 22  | 
|
| 41447 | 23  | 
val indices: prop_formula -> int list (* set of all variable indices *)  | 
24  | 
val maxidx: prop_formula -> int (* maximal variable index *)  | 
|
| 14452 | 25  | 
|
| 41447 | 26  | 
val exists: prop_formula list -> prop_formula (* finite disjunction *)  | 
27  | 
val all: prop_formula list -> prop_formula (* finite conjunction *)  | 
|
28  | 
val dot_product: prop_formula list * prop_formula list -> prop_formula  | 
|
| 14452 | 29  | 
|
| 41447 | 30  | 
val is_nnf: prop_formula -> bool (* returns true iff the formula is in negation normal form *)  | 
31  | 
val is_cnf: prop_formula -> bool (* returns true iff the formula is in conjunctive normal form *)  | 
|
| 
22441
 
7da872d34ace
is_nnf/is_cnf added, defcnf does nothing now if the formula already is in CNF
 
webertj 
parents: 
20442 
diff
changeset
 | 
32  | 
|
| 41447 | 33  | 
val nnf: prop_formula -> prop_formula (* negation normal form *)  | 
34  | 
val cnf: prop_formula -> prop_formula (* conjunctive normal form *)  | 
|
35  | 
val defcnf: prop_formula -> prop_formula (* definitional cnf *)  | 
|
| 
16909
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
36  | 
|
| 41447 | 37  | 
val eval: (int -> bool) -> prop_formula -> bool (* semantics *)  | 
| 
17809
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
38  | 
|
| 41447 | 39  | 
(* propositional representation of HOL terms *)  | 
40  | 
val prop_formula_of_term: term -> int Termtab.table -> prop_formula * int Termtab.table  | 
|
41  | 
(* HOL term representation of propositional formulae *)  | 
|
42  | 
val term_of_prop_formula: prop_formula -> term  | 
|
| 14452 | 43  | 
end;  | 
44  | 
||
45  | 
structure PropLogic : PROP_LOGIC =  | 
|
46  | 
struct  | 
|
47  | 
||
48  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 14753 | 49  | 
(* prop_formula: formulas of propositional logic, built from Boolean *)  | 
| 14452 | 50  | 
(* variables (referred to by index) and True/False using *)  | 
51  | 
(* not/or/and *)  | 
|
52  | 
(* ------------------------------------------------------------------------- *)  | 
|
53  | 
||
| 41447 | 54  | 
datatype prop_formula =  | 
55  | 
True  | 
|
56  | 
| False  | 
|
57  | 
| BoolVar of int (* NOTE: only use indices >= 1 *)  | 
|
58  | 
| Not of prop_formula  | 
|
59  | 
| Or of prop_formula * prop_formula  | 
|
60  | 
| And of prop_formula * prop_formula;  | 
|
| 14452 | 61  | 
|
62  | 
(* ------------------------------------------------------------------------- *)  | 
|
63  | 
(* The following constructor functions make sure that True and False do not *)  | 
|
64  | 
(* occur within any of the other connectives (i.e. Not, Or, And), and *)  | 
|
65  | 
(* perform double-negation elimination. *)  | 
|
66  | 
(* ------------------------------------------------------------------------- *)  | 
|
67  | 
||
| 41447 | 68  | 
fun SNot True = False  | 
69  | 
| SNot False = True  | 
|
70  | 
| SNot (Not fm) = fm  | 
|
71  | 
| SNot fm = Not fm;  | 
|
| 14452 | 72  | 
|
| 41447 | 73  | 
fun SOr (True, _) = True  | 
74  | 
| SOr (_, True) = True  | 
|
75  | 
| SOr (False, fm) = fm  | 
|
76  | 
| SOr (fm, False) = fm  | 
|
77  | 
| SOr (fm1, fm2) = Or (fm1, fm2);  | 
|
| 14452 | 78  | 
|
| 41447 | 79  | 
fun SAnd (True, fm) = fm  | 
80  | 
| SAnd (fm, True) = fm  | 
|
81  | 
| SAnd (False, _) = False  | 
|
82  | 
| SAnd (_, False) = False  | 
|
83  | 
| SAnd (fm1, fm2) = And (fm1, fm2);  | 
|
| 14452 | 84  | 
|
85  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 
16909
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
86  | 
(* simplify: eliminates True/False below other connectives, and double- *)  | 
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
87  | 
(* negation *)  | 
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
88  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
89  | 
|
| 41447 | 90  | 
fun simplify (Not fm) = SNot (simplify fm)  | 
91  | 
| simplify (Or (fm1, fm2)) = SOr (simplify fm1, simplify fm2)  | 
|
92  | 
| simplify (And (fm1, fm2)) = SAnd (simplify fm1, simplify fm2)  | 
|
93  | 
| simplify fm = fm;  | 
|
| 
16909
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
94  | 
|
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
95  | 
(* ------------------------------------------------------------------------- *)  | 
| 14753 | 96  | 
(* indices: collects all indices of Boolean variables that occur in a *)  | 
| 14452 | 97  | 
(* propositional formula 'fm'; no duplicates *)  | 
98  | 
(* ------------------------------------------------------------------------- *)  | 
|
99  | 
||
| 41447 | 100  | 
fun indices True = []  | 
101  | 
| indices False = []  | 
|
102  | 
| indices (BoolVar i) = [i]  | 
|
103  | 
| indices (Not fm) = indices fm  | 
|
104  | 
| indices (Or (fm1, fm2)) = union (op =) (indices fm1) (indices fm2)  | 
|
105  | 
| indices (And (fm1, fm2)) = union (op =) (indices fm1) (indices fm2);  | 
|
| 14452 | 106  | 
|
107  | 
(* ------------------------------------------------------------------------- *)  | 
|
108  | 
(* maxidx: computes the maximal variable index occuring in a formula of *)  | 
|
109  | 
(* propositional logic 'fm'; 0 if 'fm' contains no variable *)  | 
|
110  | 
(* ------------------------------------------------------------------------- *)  | 
|
111  | 
||
| 41447 | 112  | 
fun maxidx True = 0  | 
113  | 
| maxidx False = 0  | 
|
114  | 
| maxidx (BoolVar i) = i  | 
|
115  | 
| maxidx (Not fm) = maxidx fm  | 
|
116  | 
| maxidx (Or (fm1, fm2)) = Int.max (maxidx fm1, maxidx fm2)  | 
|
117  | 
| maxidx (And (fm1, fm2)) = Int.max (maxidx fm1, maxidx fm2);  | 
|
| 
16909
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
118  | 
|
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
119  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
120  | 
(* exists: computes the disjunction over a list 'xs' of propositional *)  | 
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
121  | 
(* formulas *)  | 
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
122  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
123  | 
|
| 41447 | 124  | 
fun exists xs = Library.foldl SOr (False, xs);  | 
| 
16909
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
125  | 
|
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
126  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
127  | 
(* all: computes the conjunction over a list 'xs' of propositional formulas *)  | 
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
128  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
129  | 
|
| 41447 | 130  | 
fun all xs = Library.foldl SAnd (True, xs);  | 
| 
16909
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
131  | 
|
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
132  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
133  | 
(* dot_product: ([x1,...,xn], [y1,...,yn]) -> x1*y1+...+xn*yn *)  | 
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
134  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
acbc7a9c3864
defcnf renamed to auxcnf, new defcnf algorithm added, simplify added
 
webertj 
parents: 
16907 
diff
changeset
 | 
135  | 
|
| 41447 | 136  | 
fun dot_product (xs, ys) = exists (map SAnd (xs ~~ ys));  | 
| 14452 | 137  | 
|
138  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 31220 | 139  | 
(* is_nnf: returns 'true' iff the formula is in negation normal form (i.e., *)  | 
| 
22441
 
7da872d34ace
is_nnf/is_cnf added, defcnf does nothing now if the formula already is in CNF
 
webertj 
parents: 
20442 
diff
changeset
 | 
140  | 
(* only variables may be negated, but not subformulas). *)  | 
| 
 
7da872d34ace
is_nnf/is_cnf added, defcnf does nothing now if the formula already is in CNF
 
webertj 
parents: 
20442 
diff
changeset
 | 
141  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
7da872d34ace
is_nnf/is_cnf added, defcnf does nothing now if the formula already is in CNF
 
webertj 
parents: 
20442 
diff
changeset
 | 
142  | 
|
| 41447 | 143  | 
local  | 
144  | 
fun is_literal (BoolVar _) = true  | 
|
145  | 
| is_literal (Not (BoolVar _)) = true  | 
|
146  | 
| is_literal _ = false  | 
|
147  | 
fun is_conj_disj (Or (fm1, fm2)) = is_conj_disj fm1 andalso is_conj_disj fm2  | 
|
148  | 
| is_conj_disj (And (fm1, fm2)) = is_conj_disj fm1 andalso is_conj_disj fm2  | 
|
149  | 
| is_conj_disj fm = is_literal fm  | 
|
150  | 
in  | 
|
151  | 
fun is_nnf True = true  | 
|
152  | 
| is_nnf False = true  | 
|
153  | 
| is_nnf fm = is_conj_disj fm  | 
|
154  | 
end;  | 
|
| 
22441
 
7da872d34ace
is_nnf/is_cnf added, defcnf does nothing now if the formula already is in CNF
 
webertj 
parents: 
20442 
diff
changeset
 | 
155  | 
|
| 
 
7da872d34ace
is_nnf/is_cnf added, defcnf does nothing now if the formula already is in CNF
 
webertj 
parents: 
20442 
diff
changeset
 | 
156  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
7da872d34ace
is_nnf/is_cnf added, defcnf does nothing now if the formula already is in CNF
 
webertj 
parents: 
20442 
diff
changeset
 | 
157  | 
(* is_cnf: returns 'true' iff the formula is in conjunctive normal form *)  | 
| 31220 | 158  | 
(* (i.e., a conjunction of disjunctions of literals). 'is_cnf' *)  | 
159  | 
(* implies 'is_nnf'. *)  | 
|
| 
22441
 
7da872d34ace
is_nnf/is_cnf added, defcnf does nothing now if the formula already is in CNF
 
webertj 
parents: 
20442 
diff
changeset
 | 
160  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
7da872d34ace
is_nnf/is_cnf added, defcnf does nothing now if the formula already is in CNF
 
webertj 
parents: 
20442 
diff
changeset
 | 
161  | 
|
| 41447 | 162  | 
local  | 
163  | 
fun is_literal (BoolVar _) = true  | 
|
164  | 
| is_literal (Not (BoolVar _)) = true  | 
|
165  | 
| is_literal _ = false  | 
|
166  | 
fun is_disj (Or (fm1, fm2)) = is_disj fm1 andalso is_disj fm2  | 
|
167  | 
| is_disj fm = is_literal fm  | 
|
168  | 
fun is_conj (And (fm1, fm2)) = is_conj fm1 andalso is_conj fm2  | 
|
169  | 
| is_conj fm = is_disj fm  | 
|
170  | 
in  | 
|
171  | 
fun is_cnf True = true  | 
|
172  | 
| is_cnf False = true  | 
|
173  | 
| is_cnf fm = is_conj fm  | 
|
174  | 
end;  | 
|
| 
22441
 
7da872d34ace
is_nnf/is_cnf added, defcnf does nothing now if the formula already is in CNF
 
webertj 
parents: 
20442 
diff
changeset
 | 
175  | 
|
| 
 
7da872d34ace
is_nnf/is_cnf added, defcnf does nothing now if the formula already is in CNF
 
webertj 
parents: 
20442 
diff
changeset
 | 
176  | 
(* ------------------------------------------------------------------------- *)  | 
| 14452 | 177  | 
(* nnf: computes the negation normal form of a formula 'fm' of propositional *)  | 
| 31220 | 178  | 
(* logic (i.e., only variables may be negated, but not subformulas). *)  | 
179  | 
(* Simplification (cf. 'simplify') is performed as well. Not *)  | 
|
180  | 
(* surprisingly, 'is_nnf o nnf' always returns 'true'. 'nnf fm' returns *)  | 
|
181  | 
(* 'fm' if (and only if) 'is_nnf fm' returns 'true'. *)  | 
|
| 14452 | 182  | 
(* ------------------------------------------------------------------------- *)  | 
183  | 
||
| 41447 | 184  | 
fun nnf fm =  | 
185  | 
let  | 
|
186  | 
fun  | 
|
187  | 
(* constants *)  | 
|
188  | 
nnf_aux True = True  | 
|
189  | 
| nnf_aux False = False  | 
|
190  | 
(* variables *)  | 
|
191  | 
| nnf_aux (BoolVar i) = (BoolVar i)  | 
|
192  | 
(* 'or' and 'and' as outermost connectives are left untouched *)  | 
|
193  | 
| nnf_aux (Or (fm1, fm2)) = SOr (nnf_aux fm1, nnf_aux fm2)  | 
|
194  | 
| nnf_aux (And (fm1, fm2)) = SAnd (nnf_aux fm1, nnf_aux fm2)  | 
|
195  | 
(* 'not' + constant *)  | 
|
196  | 
| nnf_aux (Not True) = False  | 
|
197  | 
| nnf_aux (Not False) = True  | 
|
198  | 
(* 'not' + variable *)  | 
|
199  | 
| nnf_aux (Not (BoolVar i)) = Not (BoolVar i)  | 
|
200  | 
(* pushing 'not' inside of 'or'/'and' using de Morgan's laws *)  | 
|
201  | 
| nnf_aux (Not (Or (fm1, fm2))) = SAnd (nnf_aux (SNot fm1), nnf_aux (SNot fm2))  | 
|
202  | 
| nnf_aux (Not (And (fm1, fm2))) = SOr (nnf_aux (SNot fm1), nnf_aux (SNot fm2))  | 
|
203  | 
(* double-negation elimination *)  | 
|
204  | 
| nnf_aux (Not (Not fm)) = nnf_aux fm  | 
|
205  | 
in  | 
|
206  | 
if is_nnf fm then fm  | 
|
207  | 
else nnf_aux fm  | 
|
208  | 
end;  | 
|
| 14452 | 209  | 
|
210  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 31220 | 211  | 
(* cnf: computes the conjunctive normal form (i.e., a conjunction of *)  | 
212  | 
(* disjunctions of literals) of a formula 'fm' of propositional logic. *)  | 
|
213  | 
(* Simplification (cf. 'simplify') is performed as well. The result *)  | 
|
214  | 
(* is equivalent to 'fm', but may be exponentially longer. Not *)  | 
|
215  | 
(* surprisingly, 'is_cnf o cnf' always returns 'true'. 'cnf fm' returns *)  | 
|
216  | 
(* 'fm' if (and only if) 'is_cnf fm' returns 'true'. *)  | 
|
| 14452 | 217  | 
(* ------------------------------------------------------------------------- *)  | 
218  | 
||
| 41447 | 219  | 
fun cnf fm =  | 
220  | 
let  | 
|
221  | 
(* function to push an 'Or' below 'And's, using distributive laws *)  | 
|
222  | 
fun cnf_or (And (fm11, fm12), fm2) =  | 
|
223  | 
And (cnf_or (fm11, fm2), cnf_or (fm12, fm2))  | 
|
224  | 
| cnf_or (fm1, And (fm21, fm22)) =  | 
|
225  | 
And (cnf_or (fm1, fm21), cnf_or (fm1, fm22))  | 
|
226  | 
(* neither subformula contains 'And' *)  | 
|
227  | 
| cnf_or (fm1, fm2) = Or (fm1, fm2)  | 
|
228  | 
fun cnf_from_nnf True = True  | 
|
229  | 
| cnf_from_nnf False = False  | 
|
230  | 
| cnf_from_nnf (BoolVar i) = BoolVar i  | 
|
231  | 
(* 'fm' must be a variable since the formula is in NNF *)  | 
|
232  | 
| cnf_from_nnf (Not fm) = Not fm  | 
|
233  | 
(* 'Or' may need to be pushed below 'And' *)  | 
|
234  | 
| cnf_from_nnf (Or (fm1, fm2)) =  | 
|
235  | 
cnf_or (cnf_from_nnf fm1, cnf_from_nnf fm2)  | 
|
236  | 
(* 'And' as outermost connective is left untouched *)  | 
|
237  | 
| cnf_from_nnf (And (fm1, fm2)) =  | 
|
238  | 
And (cnf_from_nnf fm1, cnf_from_nnf fm2)  | 
|
239  | 
in  | 
|
240  | 
if is_cnf fm then fm  | 
|
241  | 
else (cnf_from_nnf o nnf) fm  | 
|
242  | 
end;  | 
|
| 14452 | 243  | 
|
244  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 31220 | 245  | 
(* defcnf: computes a definitional conjunctive normal form of a formula 'fm' *)  | 
246  | 
(* of propositional logic. Simplification (cf. 'simplify') is performed *)  | 
|
247  | 
(* as well. 'defcnf' may introduce auxiliary Boolean variables to avoid *)  | 
|
248  | 
(* an exponential blowup of the formula. The result is equisatisfiable *)  | 
|
249  | 
(* (i.e., satisfiable if and only if 'fm' is satisfiable), but not *)  | 
|
250  | 
(* necessarily equivalent to 'fm'. Not surprisingly, 'is_cnf o defcnf' *)  | 
|
251  | 
(* always returns 'true'. 'defcnf fm' returns 'fm' if (and only if) *)  | 
|
252  | 
(* 'is_cnf fm' returns 'true'. *)  | 
|
| 14452 | 253  | 
(* ------------------------------------------------------------------------- *)  | 
254  | 
||
| 41447 | 255  | 
fun defcnf fm =  | 
256  | 
if is_cnf fm then fm  | 
|
257  | 
else  | 
|
258  | 
let  | 
|
259  | 
val fm' = nnf fm  | 
|
260  | 
(* 'new' specifies the next index that is available to introduce an auxiliary variable *)  | 
|
261  | 
(* int ref *)  | 
|
262  | 
val new = Unsynchronized.ref (maxidx fm' + 1)  | 
|
263  | 
(* unit -> int *)  | 
|
264  | 
fun new_idx () = let val idx = !new in new := idx+1; idx end  | 
|
265  | 
(* replaces 'And' by an auxiliary variable (and its definition) *)  | 
|
266  | 
(* prop_formula -> prop_formula * prop_formula list *)  | 
|
267  | 
fun defcnf_or (And x) =  | 
|
268  | 
let  | 
|
269  | 
val i = new_idx ()  | 
|
270  | 
in  | 
|
271  | 
(* Note that definitions are in NNF, but not CNF. *)  | 
|
272  | 
(BoolVar i, [Or (Not (BoolVar i), And x)])  | 
|
273  | 
end  | 
|
274  | 
| defcnf_or (Or (fm1, fm2)) =  | 
|
275  | 
let  | 
|
276  | 
val (fm1', defs1) = defcnf_or fm1  | 
|
277  | 
val (fm2', defs2) = defcnf_or fm2  | 
|
278  | 
in  | 
|
279  | 
(Or (fm1', fm2'), defs1 @ defs2)  | 
|
280  | 
end  | 
|
281  | 
| defcnf_or fm = (fm, [])  | 
|
282  | 
(* prop_formula -> prop_formula *)  | 
|
283  | 
fun defcnf_from_nnf True = True  | 
|
284  | 
| defcnf_from_nnf False = False  | 
|
285  | 
| defcnf_from_nnf (BoolVar i) = BoolVar i  | 
|
286  | 
(* 'fm' must be a variable since the formula is in NNF *)  | 
|
287  | 
| defcnf_from_nnf (Not fm) = Not fm  | 
|
288  | 
(* 'Or' may need to be pushed below 'And' *)  | 
|
289  | 
(* 'Or' of literal and 'And': use distributivity *)  | 
|
290  | 
| defcnf_from_nnf (Or (BoolVar i, And (fm1, fm2))) =  | 
|
291  | 
And (defcnf_from_nnf (Or (BoolVar i, fm1)),  | 
|
292  | 
defcnf_from_nnf (Or (BoolVar i, fm2)))  | 
|
293  | 
| defcnf_from_nnf (Or (Not (BoolVar i), And (fm1, fm2))) =  | 
|
294  | 
And (defcnf_from_nnf (Or (Not (BoolVar i), fm1)),  | 
|
295  | 
defcnf_from_nnf (Or (Not (BoolVar i), fm2)))  | 
|
296  | 
| defcnf_from_nnf (Or (And (fm1, fm2), BoolVar i)) =  | 
|
297  | 
And (defcnf_from_nnf (Or (fm1, BoolVar i)),  | 
|
298  | 
defcnf_from_nnf (Or (fm2, BoolVar i)))  | 
|
299  | 
| defcnf_from_nnf (Or (And (fm1, fm2), Not (BoolVar i))) =  | 
|
300  | 
And (defcnf_from_nnf (Or (fm1, Not (BoolVar i))),  | 
|
301  | 
defcnf_from_nnf (Or (fm2, Not (BoolVar i))))  | 
|
302  | 
(* all other cases: turn the formula into a disjunction of literals, *)  | 
|
303  | 
(* adding definitions as necessary *)  | 
|
304  | 
| defcnf_from_nnf (Or x) =  | 
|
305  | 
let  | 
|
306  | 
val (fm, defs) = defcnf_or (Or x)  | 
|
307  | 
val cnf_defs = map defcnf_from_nnf defs  | 
|
308  | 
in  | 
|
309  | 
all (fm :: cnf_defs)  | 
|
310  | 
end  | 
|
311  | 
(* 'And' as outermost connective is left untouched *)  | 
|
312  | 
| defcnf_from_nnf (And (fm1, fm2)) =  | 
|
313  | 
And (defcnf_from_nnf fm1, defcnf_from_nnf fm2)  | 
|
314  | 
in  | 
|
315  | 
defcnf_from_nnf fm'  | 
|
316  | 
end;  | 
|
| 14452 | 317  | 
|
318  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 14753 | 319  | 
(* eval: given an assignment 'a' of Boolean values to variable indices, the *)  | 
| 14452 | 320  | 
(* truth value of a propositional formula 'fm' is computed *)  | 
321  | 
(* ------------------------------------------------------------------------- *)  | 
|
322  | 
||
| 41447 | 323  | 
fun eval a True = true  | 
324  | 
| eval a False = false  | 
|
325  | 
| eval a (BoolVar i) = (a i)  | 
|
326  | 
| eval a (Not fm) = not (eval a fm)  | 
|
327  | 
| eval a (Or (fm1, fm2)) = (eval a fm1) orelse (eval a fm2)  | 
|
328  | 
| eval a (And (fm1, fm2)) = (eval a fm1) andalso (eval a fm2);  | 
|
| 14452 | 329  | 
|
| 
17809
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
330  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
331  | 
(* prop_formula_of_term: returns the propositional structure of a HOL term, *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
332  | 
(* with subterms replaced by Boolean variables. Also returns a table *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
333  | 
(* of terms and corresponding variables that extends the table that was *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
334  | 
(* given as an argument. Usually, you'll just want to use *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
335  | 
(* 'Termtab.empty' as value for 'table'. *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
336  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
337  | 
|
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
338  | 
(* Note: The implementation is somewhat optimized; the next index to be used *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
339  | 
(* is computed only when it is actually needed. However, when *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
340  | 
(* 'prop_formula_of_term' is invoked many times, it might be more *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
341  | 
(* efficient to pass and return this value as an additional parameter, *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
342  | 
(* so that it does not have to be recomputed (by folding over the *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
343  | 
(* table) for each invocation. *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
344  | 
|
| 41447 | 345  | 
fun prop_formula_of_term t table =  | 
346  | 
let  | 
|
347  | 
val next_idx_is_valid = Unsynchronized.ref false  | 
|
348  | 
val next_idx = Unsynchronized.ref 0  | 
|
349  | 
fun get_next_idx () =  | 
|
350  | 
if !next_idx_is_valid then  | 
|
351  | 
Unsynchronized.inc next_idx  | 
|
352  | 
else (  | 
|
353  | 
next_idx := Termtab.fold (Integer.max o snd) table 0;  | 
|
354  | 
next_idx_is_valid := true;  | 
|
355  | 
Unsynchronized.inc next_idx  | 
|
356  | 
)  | 
|
357  | 
    fun aux (Const (@{const_name True}, _)) table = (True, table)
 | 
|
358  | 
      | aux (Const (@{const_name False}, _)) table = (False, table)
 | 
|
359  | 
      | aux (Const (@{const_name Not}, _) $ x) table = apfst Not (aux x table)
 | 
|
360  | 
      | aux (Const (@{const_name HOL.disj}, _) $ x $ y) table =
 | 
|
361  | 
let  | 
|
362  | 
val (fm1, table1) = aux x table  | 
|
363  | 
val (fm2, table2) = aux y table1  | 
|
364  | 
in  | 
|
365  | 
(Or (fm1, fm2), table2)  | 
|
366  | 
end  | 
|
367  | 
      | aux (Const (@{const_name HOL.conj}, _) $ x $ y) table =
 | 
|
368  | 
let  | 
|
369  | 
val (fm1, table1) = aux x table  | 
|
370  | 
val (fm2, table2) = aux y table1  | 
|
371  | 
in  | 
|
372  | 
(And (fm1, fm2), table2)  | 
|
373  | 
end  | 
|
374  | 
| aux x table =  | 
|
375  | 
(case Termtab.lookup table x of  | 
|
376  | 
SOME i => (BoolVar i, table)  | 
|
377  | 
| NONE =>  | 
|
378  | 
let  | 
|
379  | 
val i = get_next_idx ()  | 
|
380  | 
in  | 
|
381  | 
(BoolVar i, Termtab.update (x, i) table)  | 
|
382  | 
end)  | 
|
383  | 
in  | 
|
384  | 
aux t table  | 
|
385  | 
end;  | 
|
| 
17809
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
16913 
diff
changeset
 | 
386  | 
|
| 20442 | 387  | 
(* ------------------------------------------------------------------------- *)  | 
388  | 
(* term_of_prop_formula: returns a HOL term that corresponds to a *)  | 
|
389  | 
(* propositional formula, with Boolean variables replaced by Free's *)  | 
|
390  | 
(* ------------------------------------------------------------------------- *)  | 
|
391  | 
||
392  | 
(* Note: A more generic implementation should take another argument of type *)  | 
|
393  | 
(* Term.term Inttab.table (or so) that specifies HOL terms for some *)  | 
|
394  | 
(* Boolean variables in the formula, similar to 'prop_formula_of_term' *)  | 
|
395  | 
(* (but the other way round). *)  | 
|
396  | 
||
| 41447 | 397  | 
fun term_of_prop_formula True = HOLogic.true_const  | 
398  | 
| term_of_prop_formula False = HOLogic.false_const  | 
|
399  | 
  | term_of_prop_formula (BoolVar i) = Free ("v" ^ Int.toString i, HOLogic.boolT)
 | 
|
400  | 
| term_of_prop_formula (Not fm) = HOLogic.mk_not (term_of_prop_formula fm)  | 
|
401  | 
| term_of_prop_formula (Or (fm1, fm2)) =  | 
|
402  | 
HOLogic.mk_disj (term_of_prop_formula fm1, term_of_prop_formula fm2)  | 
|
403  | 
| term_of_prop_formula (And (fm1, fm2)) =  | 
|
404  | 
HOLogic.mk_conj (term_of_prop_formula fm1, term_of_prop_formula fm2);  | 
|
| 20442 | 405  | 
|
| 14452 | 406  | 
end;  |