| author | wenzelm | 
| Thu, 05 Jan 2012 14:48:41 +0100 | |
| changeset 46123 | aa5c367ee579 | 
| parent 45204 | 5e4a1270c000 | 
| child 46708 | b138dee7bed3 | 
| permissions | -rw-r--r-- | 
| 24127 | 1  | 
|
2  | 
header {* Meson test cases *}
 | 
|
3  | 
||
4  | 
theory Meson_Test  | 
|
5  | 
imports Main  | 
|
6  | 
begin  | 
|
7  | 
||
8  | 
text {*
 | 
|
9  | 
WARNING: there are many potential conflicts between variables used  | 
|
10  | 
below and constants declared in HOL!  | 
|
11  | 
*}  | 
|
12  | 
||
| 
45204
 
5e4a1270c000
hide typedef-generated constants Product_Type.prod and Sum_Type.sum
 
huffman 
parents: 
43965 
diff
changeset
 | 
13  | 
hide_const (open) implies union inter subset quotient  | 
| 24127 | 14  | 
|
15  | 
text {*
 | 
|
16  | 
Test data for the MESON proof procedure  | 
|
17  | 
(Excludes the equality problems 51, 52, 56, 58)  | 
|
18  | 
*}  | 
|
19  | 
||
20  | 
||
21  | 
subsection {* Interactive examples *}
 | 
|
22  | 
||
| 
37777
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
23  | 
lemma problem_25:  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
24  | 
"(\<exists>x. P x) & (\<forall>x. L x --> ~ (M x & R x)) & (\<forall>x. P x --> (M x & L x)) & ((\<forall>x. P x --> Q x) | (\<exists>x. P x & R x)) --> (\<exists>x. Q x & P x)"  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
25  | 
apply (rule ccontr)  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
26  | 
  ML_prf {*
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
27  | 
    val prem25 = Thm.assume @{cprop "\<not> ?thesis"};
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
28  | 
    val nnf25 = Meson.make_nnf @{context} prem25;
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
29  | 
    val xsko25 = Meson.skolemize @{context} nnf25;
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
30  | 
*}  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
31  | 
  apply (tactic {* cut_facts_tac [xsko25] 1 THEN REPEAT (etac exE 1) *})
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
32  | 
  ML_val {*
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
33  | 
    val [_, sko25] = #prems (#1 (Subgoal.focus @{context} 1 (#goal @{Isar.goal})));
 | 
| 43965 | 34  | 
    val clauses25 = Meson.make_clauses @{context} [sko25];   (*7 clauses*)
 | 
| 
37777
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
35  | 
val horns25 = Meson.make_horns clauses25; (*16 Horn clauses*)  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
36  | 
val go25 :: _ = Meson.gocls clauses25;  | 
| 24127 | 37  | 
|
| 
37777
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
38  | 
    Goal.prove @{context} [] [] @{prop False} (fn _ =>
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
39  | 
rtac go25 1 THEN  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
40  | 
Meson.depth_prolog_tac horns25);  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
41  | 
*}  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
42  | 
oops  | 
| 24127 | 43  | 
|
| 
37777
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
44  | 
lemma problem_26:  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
45  | 
"((\<exists>x. p x) = (\<exists>x. q x)) & (\<forall>x. \<forall>y. p x & q y --> (r x = s y)) --> ((\<forall>x. p x --> r x) = (\<forall>x. q x --> s x))"  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
46  | 
apply (rule ccontr)  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
47  | 
  ML_prf {*
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
48  | 
    val prem26 = Thm.assume @{cprop "\<not> ?thesis"}
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
49  | 
    val nnf26 = Meson.make_nnf @{context} prem26;
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
50  | 
    val xsko26 = Meson.skolemize @{context} nnf26;
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
51  | 
*}  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
52  | 
  apply (tactic {* cut_facts_tac [xsko26] 1 THEN REPEAT (etac exE 1) *})
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
53  | 
  ML_val {*
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
54  | 
    val [_, sko26] = #prems (#1 (Subgoal.focus @{context} 1 (#goal @{Isar.goal})));
 | 
| 43965 | 55  | 
    val clauses26 = Meson.make_clauses @{context} [sko26];                   (*9 clauses*)
 | 
| 
37777
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
56  | 
val horns26 = Meson.make_horns clauses26; (*24 Horn clauses*)  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
57  | 
val go26 :: _ = Meson.gocls clauses26;  | 
| 24127 | 58  | 
|
| 
37777
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
59  | 
    Goal.prove @{context} [] [] @{prop False} (fn _ =>
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
60  | 
rtac go26 1 THEN  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
61  | 
Meson.depth_prolog_tac horns26); (*7 ms*)  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
62  | 
(*Proof is of length 107!!*)  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
63  | 
*}  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
64  | 
oops  | 
| 24127 | 65  | 
|
| 
37777
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
66  | 
lemma problem_43: -- "NOW PROVED AUTOMATICALLY!!" (*16 Horn clauses*)  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
67  | 
"(\<forall>x. \<forall>y. q x y = (\<forall>z. p z x = (p z y::bool))) --> (\<forall>x. (\<forall>y. q x y = (q y x::bool)))"  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
68  | 
apply (rule ccontr)  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
69  | 
  ML_prf {*
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
70  | 
    val prem43 = Thm.assume @{cprop "\<not> ?thesis"};
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
71  | 
    val nnf43 = Meson.make_nnf @{context} prem43;
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
72  | 
    val xsko43 = Meson.skolemize @{context} nnf43;
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
73  | 
*}  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
74  | 
  apply (tactic {* cut_facts_tac [xsko43] 1 THEN REPEAT (etac exE 1) *})
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
75  | 
  ML_val {*
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
76  | 
    val [_, sko43] = #prems (#1 (Subgoal.focus @{context} 1 (#goal @{Isar.goal})));
 | 
| 43965 | 77  | 
    val clauses43 = Meson.make_clauses @{context} [sko43];   (*6*)
 | 
| 
37777
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
78  | 
val horns43 = Meson.make_horns clauses43; (*16*)  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
79  | 
val go43 :: _ = Meson.gocls clauses43;  | 
| 24127 | 80  | 
|
| 
37777
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
81  | 
    Goal.prove @{context} [] [] @{prop False} (fn _ =>
 | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
82  | 
rtac go43 1 THEN  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
83  | 
Meson.best_prolog_tac Meson.size_of_subgoals horns43); (*7ms*)  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
84  | 
*}  | 
| 
 
22107b894e5a
some modernization of really ancient Meson experiments;
 
wenzelm 
parents: 
37595 
diff
changeset
 | 
85  | 
oops  | 
| 24127 | 86  | 
|
| 24128 | 87  | 
(*  | 
| 24127 | 88  | 
#1 (q x xa ==> ~ q x xa) ==> q xa x  | 
89  | 
#2 (q xa x ==> ~ q xa x) ==> q x xa  | 
|
90  | 
#3 (~ q x xa ==> q x xa) ==> ~ q xa x  | 
|
91  | 
#4 (~ q xa x ==> q xa x) ==> ~ q x xa  | 
|
92  | 
#5 [| ~ q ?U ?V ==> q ?U ?V; ~ p ?W ?U ==> p ?W ?U |] ==> p ?W ?V  | 
|
93  | 
#6 [| ~ p ?W ?U ==> p ?W ?U; p ?W ?V ==> ~ p ?W ?V |] ==> ~ q ?U ?V  | 
|
94  | 
#7 [| p ?W ?V ==> ~ p ?W ?V; ~ q ?U ?V ==> q ?U ?V |] ==> ~ p ?W ?U  | 
|
95  | 
#8 [| ~ q ?U ?V ==> q ?U ?V; ~ p ?W ?V ==> p ?W ?V |] ==> p ?W ?U  | 
|
96  | 
#9 [| ~ p ?W ?V ==> p ?W ?V; p ?W ?U ==> ~ p ?W ?U |] ==> ~ q ?U ?V  | 
|
97  | 
#10 [| p ?W ?U ==> ~ p ?W ?U; ~ q ?U ?V ==> q ?U ?V |] ==> ~ p ?W ?V  | 
|
98  | 
#11 [| p (xb ?U ?V) ?U ==> ~ p (xb ?U ?V) ?U;  | 
|
99  | 
p (xb ?U ?V) ?V ==> ~ p (xb ?U ?V) ?V |] ==> q ?U ?V  | 
|
100  | 
#12 [| p (xb ?U ?V) ?V ==> ~ p (xb ?U ?V) ?V; q ?U ?V ==> ~ q ?U ?V |] ==>  | 
|
101  | 
p (xb ?U ?V) ?U  | 
|
102  | 
#13 [| q ?U ?V ==> ~ q ?U ?V; p (xb ?U ?V) ?U ==> ~ p (xb ?U ?V) ?U |] ==>  | 
|
103  | 
p (xb ?U ?V) ?V  | 
|
104  | 
#14 [| ~ p (xb ?U ?V) ?U ==> p (xb ?U ?V) ?U;  | 
|
105  | 
~ p (xb ?U ?V) ?V ==> p (xb ?U ?V) ?V |] ==> q ?U ?V  | 
|
106  | 
#15 [| ~ p (xb ?U ?V) ?V ==> p (xb ?U ?V) ?V; q ?U ?V ==> ~ q ?U ?V |] ==>  | 
|
107  | 
~ p (xb ?U ?V) ?U  | 
|
108  | 
#16 [| q ?U ?V ==> ~ q ?U ?V; ~ p (xb ?U ?V) ?U ==> p (xb ?U ?V) ?U |] ==>  | 
|
109  | 
~ p (xb ?U ?V) ?V  | 
|
110  | 
||
111  | 
And here is the proof! (Unkn is the start state after use of goal clause)  | 
|
112  | 
[Unkn, Res ([Thm "#14"], false, 1), Res ([Thm "#5"], false, 1),  | 
|
113  | 
Res ([Thm "#1"], false, 1), Asm 1, Res ([Thm "#13"], false, 1), Asm 2,  | 
|
114  | 
Asm 1, Res ([Thm "#13"], false, 1), Asm 1, Res ([Thm "#10"], false, 1),  | 
|
115  | 
Res ([Thm "#16"], false, 1), Asm 2, Asm 1, Res ([Thm "#1"], false, 1),  | 
|
116  | 
Asm 1, Res ([Thm "#14"], false, 1), Res ([Thm "#5"], false, 1),  | 
|
117  | 
Res ([Thm "#2"], false, 1), Asm 1, Res ([Thm "#13"], false, 1), Asm 2,  | 
|
118  | 
Asm 1, Res ([Thm "#8"], false, 1), Res ([Thm "#2"], false, 1), Asm 1,  | 
|
119  | 
Res ([Thm "#12"], false, 1), Asm 2, Asm 1] : lderiv list  | 
|
120  | 
*)  | 
|
121  | 
||
122  | 
||
123  | 
text {*
 | 
|
124  | 
MORE and MUCH HARDER test data for the MESON proof procedure  | 
|
125  | 
(courtesy John Harrison).  | 
|
126  | 
*}  | 
|
127  | 
||
128  | 
(* ========================================================================= *)  | 
|
129  | 
(* 100 problems selected from the TPTP library *)  | 
|
130  | 
(* ========================================================================= *)  | 
|
131  | 
||
132  | 
(*  | 
|
133  | 
* Original timings for John Harrison's MESON_TAC.  | 
|
134  | 
* Timings below on a 600MHz Pentium III (perch)  | 
|
135  | 
* Some timings below refer to griffon, which is a dual 2.5GHz Power Mac G5.  | 
|
| 24128 | 136  | 
*  | 
| 24127 | 137  | 
* A few variable names have been changed to avoid clashing with constants.  | 
138  | 
*  | 
|
139  | 
* Changed numeric constants e.g. 0, 1, 2... to num0, num1, num2...  | 
|
140  | 
*  | 
|
141  | 
* Here's a list giving typical CPU times, as well as common names and  | 
|
142  | 
* literature references.  | 
|
143  | 
*  | 
|
144  | 
* BOO003-1 34.6 B2 part 1 [McCharen, et al., 1976]; Lemma proved [Overbeek, et al., 1976]; prob2_part1.ver1.in [ANL]  | 
|
145  | 
* BOO004-1 36.7 B2 part 2 [McCharen, et al., 1976]; Lemma proved [Overbeek, et al., 1976]; prob2_part2.ver1 [ANL]  | 
|
146  | 
* BOO005-1 47.4 B3 part 1 [McCharen, et al., 1976]; B5 [McCharen, et al., 1976]; Lemma proved [Overbeek, et al., 1976]; prob3_part1.ver1.in [ANL]  | 
|
147  | 
* BOO006-1 48.4 B3 part 2 [McCharen, et al., 1976]; B6 [McCharen, et al., 1976]; Lemma proved [Overbeek, et al., 1976]; prob3_part2.ver1 [ANL]  | 
|
148  | 
* BOO011-1 19.0 B7 [McCharen, et al., 1976]; prob7.ver1 [ANL]  | 
|
149  | 
* CAT001-3 45.2 C1 [McCharen, et al., 1976]; p1.ver3.in [ANL]  | 
|
150  | 
* CAT003-3 10.5 C3 [McCharen, et al., 1976]; p3.ver3.in [ANL]  | 
|
151  | 
* CAT005-1 480.1 C5 [McCharen, et al., 1976]; p5.ver1.in [ANL]  | 
|
152  | 
* CAT007-1 11.9 C7 [McCharen, et al., 1976]; p7.ver1.in [ANL]  | 
|
153  | 
* CAT018-1 81.3 p18.ver1.in [ANL]  | 
|
154  | 
* COL001-2 16.0 C1 [Wos & McCune, 1988]  | 
|
155  | 
* COL023-1 5.1 [McCune & Wos, 1988]  | 
|
156  | 
* COL032-1 15.8 [McCune & Wos, 1988]  | 
|
157  | 
* COL052-2 13.2 bird4.ver2.in [ANL]  | 
|
158  | 
* COL075-2 116.9 [Jech, 1994]  | 
|
159  | 
* COM001-1 1.7 shortburst [Wilson & Minker, 1976]  | 
|
160  | 
* COM002-1 4.4 burstall [Wilson & Minker, 1976]  | 
|
161  | 
* COM002-2 7.4  | 
|
162  | 
* COM003-2 22.1 [Brushi, 1991]  | 
|
163  | 
* COM004-1 45.1  | 
|
164  | 
* GEO003-1 71.7 T3 [McCharen, et al., 1976]; t3.ver1.in [ANL]  | 
|
165  | 
* GEO017-2 78.8 D4.1 [Quaife, 1989]  | 
|
166  | 
* GEO027-3 181.5 D10.1 [Quaife, 1989]  | 
|
167  | 
* GEO058-2 104.0 R4 [Quaife, 1989]  | 
|
168  | 
* GEO079-1 2.4 GEOMETRY THEOREM [Slagle, 1967]  | 
|
169  | 
* GRP001-1 47.8 CADE-11 Competition 1 [Overbeek, 1990]; G1 [McCharen, et al., 1976]; THEOREM 1 [Lusk & McCune, 1993]; wos10 [Wilson & Minker, 1976]; xsquared.ver1.in [ANL]; [Robinson, 1963]  | 
|
170  | 
* GRP008-1 50.4 Problem 4 [Wos, 1965]; wos4 [Wilson & Minker, 1976]  | 
|
171  | 
* GRP013-1 40.2 Problem 11 [Wos, 1965]; wos11 [Wilson & Minker, 1976]  | 
|
172  | 
* GRP037-3 43.8 Problem 17 [Wos, 1965]; wos17 [Wilson & Minker, 1976]  | 
|
173  | 
* GRP031-2 3.2 ls23 [Lawrence & Starkey, 1974]; ls23 [Wilson & Minker, 1976]  | 
|
174  | 
* GRP034-4 2.5 ls26 [Lawrence & Starkey, 1974]; ls26 [Wilson & Minker, 1976]  | 
|
175  | 
* GRP047-2 11.7 [Veroff, 1992]  | 
|
176  | 
* GRP130-1 170.6 Bennett QG8 [TPTP]; QG8 [Slaney, 1993]  | 
|
177  | 
* GRP156-1 48.7 ax_mono1c [Schulz, 1995]  | 
|
178  | 
* GRP168-1 159.1 p01a [Schulz, 1995]  | 
|
179  | 
* HEN003-3 39.9 HP3 [McCharen, et al., 1976]  | 
|
180  | 
* HEN007-2 125.7 H7 [McCharen, et al., 1976]  | 
|
181  | 
* HEN008-4 62.0 H8 [McCharen, et al., 1976]  | 
|
182  | 
* HEN009-5 136.3 H9 [McCharen, et al., 1976]; hp9.ver3.in [ANL]  | 
|
183  | 
* HEN012-3 48.5 new.ver2.in [ANL]  | 
|
184  | 
* LCL010-1 370.9 EC-73 [McCune & Wos, 1992]; ec_yq.in [OTTER]  | 
|
185  | 
* LCL077-2 51.6 morgan.two.ver1.in [ANL]  | 
|
186  | 
* LCL082-1 14.6 IC-1.1 [Wos, et al., 1990]; IC-65 [McCune & Wos, 1992]; ls2 [SETHEO]; S1 [Pfenning, 1988]  | 
|
187  | 
* LCL111-1 585.6 CADE-11 Competition 6 [Overbeek, 1990]; mv25.in [OTTER]; MV-57 [McCune & Wos, 1992]; mv.in part 2 [OTTER]; ovb6 [SETHEO]; THEOREM 6 [Lusk & McCune, 1993]  | 
|
188  | 
* LCL143-1 10.9 Lattice structure theorem 2 [Bonacina, 1991]  | 
|
189  | 
* LCL182-1 271.6 Problem 2.16 [Whitehead & Russell, 1927]  | 
|
190  | 
* LCL200-1 12.0 Problem 2.46 [Whitehead & Russell, 1927]  | 
|
191  | 
* LCL215-1 214.4 Problem 2.62 [Whitehead & Russell, 1927]; Problem 2.63 [Whitehead & Russell, 1927]  | 
|
192  | 
* LCL230-2 0.2 Pelletier 5 [Pelletier, 1986]  | 
|
193  | 
* LDA003-1 68.5 Problem 3 [Jech, 1993]  | 
|
194  | 
* MSC002-1 9.2 DBABHP [Michie, et al., 1972]; DBABHP [Wilson & Minker, 1976]  | 
|
195  | 
* MSC003-1 3.2 HASPARTS-T1 [Wilson & Minker, 1976]  | 
|
196  | 
* MSC004-1 9.3 HASPARTS-T2 [Wilson & Minker, 1976]  | 
|
197  | 
* MSC005-1 1.8 Problem 5.1 [Plaisted, 1982]  | 
|
198  | 
* MSC006-1 39.0 nonob.lop [SETHEO]  | 
|
199  | 
* NUM001-1 14.0 Chang-Lee-10a [Chang, 1970]; ls28 [Lawrence & Starkey, 1974]; ls28 [Wilson & Minker, 1976]  | 
|
200  | 
* NUM021-1 52.3 ls65 [Lawrence & Starkey, 1974]; ls65 [Wilson & Minker, 1976]  | 
|
201  | 
* NUM024-1 64.6 ls75 [Lawrence & Starkey, 1974]; ls75 [Wilson & Minker, 1976]  | 
|
202  | 
* NUM180-1 621.2 LIM2.1 [Quaife]  | 
|
203  | 
* NUM228-1 575.9 TRECDEF4 cor. [Quaife]  | 
|
204  | 
* PLA002-1 37.4 Problem 5.7 [Plaisted, 1982]  | 
|
205  | 
* PLA006-1 7.2 [Segre & Elkan, 1994]  | 
|
206  | 
* PLA017-1 484.8 [Segre & Elkan, 1994]  | 
|
207  | 
* PLA022-1 19.1 [Segre & Elkan, 1994]  | 
|
208  | 
* PLA022-2 19.7 [Segre & Elkan, 1994]  | 
|
209  | 
* PRV001-1 10.3 PV1 [McCharen, et al., 1976]  | 
|
210  | 
* PRV003-1 3.9 E2 [McCharen, et al., 1976]; v2.lop [SETHEO]  | 
|
211  | 
* PRV005-1 4.3 E4 [McCharen, et al., 1976]; v4.lop [SETHEO]  | 
|
212  | 
* PRV006-1 6.0 E5 [McCharen, et al., 1976]; v5.lop [SETHEO]  | 
|
213  | 
* PRV009-1 2.2 Hoares FIND [Bledsoe, 1977]; Problem 5.5 [Plaisted, 1982]  | 
|
214  | 
* PUZ012-1 3.5 Boxes-of-fruit [Wos, 1988]; Boxes-of-fruit [Wos, et al., 1992]; boxes.ver1.in [ANL]  | 
|
215  | 
* PUZ020-1 56.6 knightknave.in [ANL]  | 
|
216  | 
* PUZ025-1 58.4 Problem 35 [Smullyan, 1978]; tandl35.ver1.in [ANL]  | 
|
217  | 
* PUZ029-1 5.1 pigs.ver1.in [ANL]  | 
|
218  | 
* RNG001-3 82.4 EX6-T? [Wilson & Minker, 1976]; ex6.lop [SETHEO]; Example 6a [Fleisig, et al., 1974]; FEX6T1 [SPRFN]; FEX6T2 [SPRFN]  | 
|
219  | 
* RNG001-5 399.8 Problem 21 [Wos, 1965]; wos21 [Wilson & Minker, 1976]  | 
|
220  | 
* RNG011-5 8.4 CADE-11 Competition Eq-10 [Overbeek, 1990]; PROBLEM 10 [Zhang, 1993]; THEOREM EQ-10 [Lusk & McCune, 1993]  | 
|
221  | 
* RNG023-6 9.1 [Stevens, 1987]  | 
|
222  | 
* RNG028-2 9.3 PROOF III [Anantharaman & Hsiang, 1990]  | 
|
223  | 
* RNG038-2 16.2 Problem 27 [Wos, 1965]; wos27 [Wilson & Minker, 1976]  | 
|
224  | 
* RNG040-2 180.5 Problem 29 [Wos, 1965]; wos29 [Wilson & Minker, 1976]  | 
|
225  | 
* RNG041-1 35.8 Problem 30 [Wos, 1965]; wos30 [Wilson & Minker, 1976]  | 
|
226  | 
* ROB010-1 205.0 Lemma 3.3 [Winker, 1990]; RA2 [Lusk & Wos, 1992]  | 
|
227  | 
* ROB013-1 23.6 Lemma 3.5 [Winker, 1990]  | 
|
228  | 
* ROB016-1 15.2 Corollary 3.7 [Winker, 1990]  | 
|
229  | 
* ROB021-1 230.4 [McCune, 1992]  | 
|
230  | 
* SET005-1 192.2 ls108 [Lawrence & Starkey, 1974]; ls108 [Wilson & Minker, 1976]  | 
|
231  | 
* SET009-1 10.5 ls116 [Lawrence & Starkey, 1974]; ls116 [Wilson & Minker, 1976]  | 
|
232  | 
* SET025-4 694.7 Lemma 10 [Boyer, et al, 1986]  | 
|
233  | 
* SET046-5 2.3 p42.in [ANL]; Pelletier 42 [Pelletier, 1986]  | 
|
234  | 
* SET047-5 3.7 p43.in [ANL]; Pelletier 43 [Pelletier, 1986]  | 
|
235  | 
* SYN034-1 2.8 QW [Michie, et al., 1972]; QW [Wilson & Minker, 1976]  | 
|
236  | 
* SYN071-1 1.9 Pelletier 48 [Pelletier, 1986]  | 
|
237  | 
* SYN349-1 61.7 Ch17N5 [Tammet, 1994]  | 
|
238  | 
* SYN352-1 5.5 Ch18N4 [Tammet, 1994]  | 
|
239  | 
* TOP001-2 61.1 Lemma 1a [Wick & McCune, 1989]  | 
|
240  | 
* TOP002-2 0.4 Lemma 1b [Wick & McCune, 1989]  | 
|
241  | 
* TOP004-1 181.6 Lemma 1d [Wick & McCune, 1989]  | 
|
242  | 
* TOP004-2 9.0 Lemma 1d [Wick & McCune, 1989]  | 
|
243  | 
* TOP005-2 139.8 Lemma 1e [Wick & McCune, 1989]  | 
|
244  | 
*)  | 
|
245  | 
||
| 24128 | 246  | 
abbreviation "EQU001_0_ax equal \<equiv> (\<forall>X. equal(X::'a,X)) &  | 
247  | 
(\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) &  | 
|
| 24127 | 248  | 
(\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z))"  | 
249  | 
||
250  | 
abbreviation "BOO002_0_ax equal INVERSE multiplicative_identity  | 
|
251  | 
additive_identity multiply product add sum \<equiv>  | 
|
| 24128 | 252  | 
(\<forall>X Y. sum(X::'a,Y,add(X::'a,Y))) &  | 
253  | 
(\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &  | 
|
254  | 
(\<forall>Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) &  | 
|
255  | 
(\<forall>Y X Z. product(X::'a,Y,Z) --> product(Y::'a,X,Z)) &  | 
|
256  | 
(\<forall>X. sum(additive_identity::'a,X,X)) &  | 
|
257  | 
(\<forall>X. sum(X::'a,additive_identity,X)) &  | 
|
258  | 
(\<forall>X. product(multiplicative_identity::'a,X,X)) &  | 
|
259  | 
(\<forall>X. product(X::'a,multiplicative_identity,X)) &  | 
|
260  | 
(\<forall>Y Z X V3 V1 V2 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & product(X::'a,V3,V4) --> sum(V1::'a,V2,V4)) &  | 
|
261  | 
(\<forall>Y Z V1 V2 X V3 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & sum(V1::'a,V2,V4) --> product(X::'a,V3,V4)) &  | 
|
262  | 
(\<forall>Y Z V3 X V1 V2 V4. product(Y::'a,X,V1) & product(Z::'a,X,V2) & sum(Y::'a,Z,V3) & product(V3::'a,X,V4) --> sum(V1::'a,V2,V4)) &  | 
|
263  | 
(\<forall>Y Z V1 V2 V3 X V4. product(Y::'a,X,V1) & product(Z::'a,X,V2) & sum(Y::'a,Z,V3) & sum(V1::'a,V2,V4) --> product(V3::'a,X,V4)) &  | 
|
264  | 
(\<forall>Y Z X V3 V1 V2 V4. sum(X::'a,Y,V1) & sum(X::'a,Z,V2) & product(Y::'a,Z,V3) & sum(X::'a,V3,V4) --> product(V1::'a,V2,V4)) &  | 
|
265  | 
(\<forall>Y Z V1 V2 X V3 V4. sum(X::'a,Y,V1) & sum(X::'a,Z,V2) & product(Y::'a,Z,V3) & product(V1::'a,V2,V4) --> sum(X::'a,V3,V4)) &  | 
|
266  | 
(\<forall>Y Z V3 X V1 V2 V4. sum(Y::'a,X,V1) & sum(Z::'a,X,V2) & product(Y::'a,Z,V3) & sum(V3::'a,X,V4) --> product(V1::'a,V2,V4)) &  | 
|
267  | 
(\<forall>Y Z V1 V2 V3 X V4. sum(Y::'a,X,V1) & sum(Z::'a,X,V2) & product(Y::'a,Z,V3) & product(V1::'a,V2,V4) --> sum(V3::'a,X,V4)) &  | 
|
268  | 
(\<forall>X. sum(INVERSE(X),X,multiplicative_identity)) &  | 
|
269  | 
(\<forall>X. sum(X::'a,INVERSE(X),multiplicative_identity)) &  | 
|
270  | 
(\<forall>X. product(INVERSE(X),X,additive_identity)) &  | 
|
271  | 
(\<forall>X. product(X::'a,INVERSE(X),additive_identity)) &  | 
|
272  | 
(\<forall>X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) &  | 
|
| 24127 | 273  | 
(\<forall>X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V))"  | 
274  | 
||
275  | 
abbreviation "BOO002_0_eq INVERSE multiply add product sum equal \<equiv>  | 
|
| 24128 | 276  | 
(\<forall>X Y W Z. equal(X::'a,Y) & sum(X::'a,W,Z) --> sum(Y::'a,W,Z)) &  | 
277  | 
(\<forall>X W Y Z. equal(X::'a,Y) & sum(W::'a,X,Z) --> sum(W::'a,Y,Z)) &  | 
|
278  | 
(\<forall>X W Z Y. equal(X::'a,Y) & sum(W::'a,Z,X) --> sum(W::'a,Z,Y)) &  | 
|
279  | 
(\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &  | 
|
280  | 
(\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &  | 
|
281  | 
(\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) &  | 
|
282  | 
(\<forall>X Y W. equal(X::'a,Y) --> equal(add(X::'a,W),add(Y::'a,W))) &  | 
|
283  | 
(\<forall>X W Y. equal(X::'a,Y) --> equal(add(W::'a,X),add(W::'a,Y))) &  | 
|
284  | 
(\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &  | 
|
285  | 
(\<forall>X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) &  | 
|
| 24127 | 286  | 
(\<forall>X Y. equal(X::'a,Y) --> equal(INVERSE(X),INVERSE(Y)))"  | 
287  | 
||
288  | 
(*51194 inferences so far. Searching to depth 13. 232.9 secs*)  | 
|
289  | 
lemma BOO003_1:  | 
|
290  | 
"EQU001_0_ax equal &  | 
|
291  | 
BOO002_0_ax equal INVERSE multiplicative_identity additive_identity multiply product add sum &  | 
|
292  | 
BOO002_0_eq INVERSE multiply add product sum equal &  | 
|
293  | 
(~product(x::'a,x,x)) --> False"  | 
|
294  | 
oops  | 
|
295  | 
||
296  | 
(*51194 inferences so far. Searching to depth 13. 204.6 secs  | 
|
297  | 
Strange! The previous problem also has 51194 inferences at depth 13. They  | 
|
298  | 
must be very similar!*)  | 
|
299  | 
lemma BOO004_1:  | 
|
300  | 
"EQU001_0_ax equal &  | 
|
301  | 
BOO002_0_ax equal INVERSE multiplicative_identity additive_identity multiply product add sum &  | 
|
302  | 
BOO002_0_eq INVERSE multiply add product sum equal &  | 
|
303  | 
(~sum(x::'a,x,x)) --> False"  | 
|
304  | 
oops  | 
|
305  | 
||
306  | 
(*74799 inferences so far. Searching to depth 13. 290.0 secs*)  | 
|
307  | 
lemma BOO005_1:  | 
|
308  | 
"EQU001_0_ax equal &  | 
|
309  | 
BOO002_0_ax equal INVERSE multiplicative_identity additive_identity multiply product add sum &  | 
|
310  | 
BOO002_0_eq INVERSE multiply add product sum equal &  | 
|
311  | 
(~sum(x::'a,multiplicative_identity,multiplicative_identity)) --> False"  | 
|
312  | 
oops  | 
|
313  | 
||
314  | 
(*74799 inferences so far. Searching to depth 13. 314.6 secs*)  | 
|
315  | 
lemma BOO006_1:  | 
|
316  | 
"EQU001_0_ax equal &  | 
|
317  | 
BOO002_0_ax equal INVERSE multiplicative_identity additive_identity multiply product add sum &  | 
|
318  | 
BOO002_0_eq INVERSE multiply add product sum equal &  | 
|
319  | 
(~product(x::'a,additive_identity,additive_identity)) --> False"  | 
|
320  | 
oops  | 
|
321  | 
||
322  | 
(*5 inferences so far. Searching to depth 5. 1.3 secs*)  | 
|
323  | 
lemma BOO011_1:  | 
|
324  | 
"EQU001_0_ax equal &  | 
|
325  | 
BOO002_0_ax equal INVERSE multiplicative_identity additive_identity multiply product add sum &  | 
|
326  | 
BOO002_0_eq INVERSE multiply add product sum equal &  | 
|
327  | 
(~equal(INVERSE(additive_identity),multiplicative_identity)) --> False"  | 
|
328  | 
by meson  | 
|
329  | 
||
330  | 
abbreviation "CAT003_0_ax f1 compos codomain domain equal there_exists equivalent \<equiv>  | 
|
| 24128 | 331  | 
(\<forall>Y X. equivalent(X::'a,Y) --> there_exists(X)) &  | 
332  | 
(\<forall>X Y. equivalent(X::'a,Y) --> equal(X::'a,Y)) &  | 
|
333  | 
(\<forall>X Y. there_exists(X) & equal(X::'a,Y) --> equivalent(X::'a,Y)) &  | 
|
334  | 
(\<forall>X. there_exists(domain(X)) --> there_exists(X)) &  | 
|
335  | 
(\<forall>X. there_exists(codomain(X)) --> there_exists(X)) &  | 
|
336  | 
(\<forall>Y X. there_exists(compos(X::'a,Y)) --> there_exists(domain(X))) &  | 
|
337  | 
(\<forall>X Y. there_exists(compos(X::'a,Y)) --> equal(domain(X),codomain(Y))) &  | 
|
338  | 
(\<forall>X Y. there_exists(domain(X)) & equal(domain(X),codomain(Y)) --> there_exists(compos(X::'a,Y))) &  | 
|
339  | 
(\<forall>X Y Z. equal(compos(X::'a,compos(Y::'a,Z)),compos(compos(X::'a,Y),Z))) &  | 
|
340  | 
(\<forall>X. equal(compos(X::'a,domain(X)),X)) &  | 
|
341  | 
(\<forall>X. equal(compos(codomain(X),X),X)) &  | 
|
342  | 
(\<forall>X Y. equivalent(X::'a,Y) --> there_exists(Y)) &  | 
|
343  | 
(\<forall>X Y. there_exists(X) & there_exists(Y) & equal(X::'a,Y) --> equivalent(X::'a,Y)) &  | 
|
344  | 
(\<forall>Y X. there_exists(compos(X::'a,Y)) --> there_exists(codomain(X))) &  | 
|
345  | 
(\<forall>X Y. there_exists(f1(X::'a,Y)) | equal(X::'a,Y)) &  | 
|
346  | 
(\<forall>X Y. equal(X::'a,f1(X::'a,Y)) | equal(Y::'a,f1(X::'a,Y)) | equal(X::'a,Y)) &  | 
|
| 24127 | 347  | 
(\<forall>X Y. equal(X::'a,f1(X::'a,Y)) & equal(Y::'a,f1(X::'a,Y)) --> equal(X::'a,Y))"  | 
348  | 
||
349  | 
abbreviation "CAT003_0_eq f1 compos codomain domain equivalent there_exists equal \<equiv>  | 
|
| 24128 | 350  | 
(\<forall>X Y. equal(X::'a,Y) & there_exists(X) --> there_exists(Y)) &  | 
351  | 
(\<forall>X Y Z. equal(X::'a,Y) & equivalent(X::'a,Z) --> equivalent(Y::'a,Z)) &  | 
|
352  | 
(\<forall>X Z Y. equal(X::'a,Y) & equivalent(Z::'a,X) --> equivalent(Z::'a,Y)) &  | 
|
353  | 
(\<forall>X Y. equal(X::'a,Y) --> equal(domain(X),domain(Y))) &  | 
|
354  | 
(\<forall>X Y. equal(X::'a,Y) --> equal(codomain(X),codomain(Y))) &  | 
|
355  | 
(\<forall>X Y Z. equal(X::'a,Y) --> equal(compos(X::'a,Z),compos(Y::'a,Z))) &  | 
|
356  | 
(\<forall>X Z Y. equal(X::'a,Y) --> equal(compos(Z::'a,X),compos(Z::'a,Y))) &  | 
|
357  | 
(\<forall>A B C. equal(A::'a,B) --> equal(f1(A::'a,C),f1(B::'a,C))) &  | 
|
| 24127 | 358  | 
(\<forall>D F' E. equal(D::'a,E) --> equal(f1(F'::'a,D),f1(F'::'a,E)))"  | 
359  | 
||
360  | 
(*4007 inferences so far. Searching to depth 9. 13 secs*)  | 
|
361  | 
lemma CAT001_3:  | 
|
362  | 
"EQU001_0_ax equal &  | 
|
363  | 
CAT003_0_ax f1 compos codomain domain equal there_exists equivalent &  | 
|
364  | 
CAT003_0_eq f1 compos codomain domain equivalent there_exists equal &  | 
|
| 24128 | 365  | 
(there_exists(compos(a::'a,b))) &  | 
366  | 
(\<forall>Y X Z. equal(compos(compos(a::'a,b),X),Y) & equal(compos(compos(a::'a,b),Z),Y) --> equal(X::'a,Z)) &  | 
|
367  | 
(there_exists(compos(b::'a,h))) &  | 
|
368  | 
(equal(compos(b::'a,h),compos(b::'a,g))) &  | 
|
| 24127 | 369  | 
(~equal(h::'a,g)) --> False"  | 
370  | 
by meson  | 
|
371  | 
||
372  | 
(*245 inferences so far. Searching to depth 7. 1.0 secs*)  | 
|
373  | 
lemma CAT003_3:  | 
|
374  | 
"EQU001_0_ax equal &  | 
|
375  | 
CAT003_0_ax f1 compos codomain domain equal there_exists equivalent &  | 
|
376  | 
CAT003_0_eq f1 compos codomain domain equivalent there_exists equal &  | 
|
| 24128 | 377  | 
(there_exists(compos(a::'a,b))) &  | 
378  | 
(\<forall>Y X Z. equal(compos(X::'a,compos(a::'a,b)),Y) & equal(compos(Z::'a,compos(a::'a,b)),Y) --> equal(X::'a,Z)) &  | 
|
379  | 
(there_exists(h)) &  | 
|
380  | 
(equal(compos(h::'a,a),compos(g::'a,a))) &  | 
|
| 24127 | 381  | 
(~equal(g::'a,h)) --> False"  | 
382  | 
by meson  | 
|
383  | 
||
384  | 
abbreviation "CAT001_0_ax equal codomain domain identity_map compos product defined \<equiv>  | 
|
| 24128 | 385  | 
(\<forall>X Y. defined(X::'a,Y) --> product(X::'a,Y,compos(X::'a,Y))) &  | 
386  | 
(\<forall>Z X Y. product(X::'a,Y,Z) --> defined(X::'a,Y)) &  | 
|
387  | 
(\<forall>X Xy Y Z. product(X::'a,Y,Xy) & defined(Xy::'a,Z) --> defined(Y::'a,Z)) &  | 
|
388  | 
(\<forall>Y Xy Z X Yz. product(X::'a,Y,Xy) & product(Y::'a,Z,Yz) & defined(Xy::'a,Z) --> defined(X::'a,Yz)) &  | 
|
389  | 
(\<forall>Xy Y Z X Yz Xyz. product(X::'a,Y,Xy) & product(Xy::'a,Z,Xyz) & product(Y::'a,Z,Yz) --> product(X::'a,Yz,Xyz)) &  | 
|
390  | 
(\<forall>Z Yz X Y. product(Y::'a,Z,Yz) & defined(X::'a,Yz) --> defined(X::'a,Y)) &  | 
|
391  | 
(\<forall>Y X Yz Xy Z. product(Y::'a,Z,Yz) & product(X::'a,Y,Xy) & defined(X::'a,Yz) --> defined(Xy::'a,Z)) &  | 
|
392  | 
(\<forall>Yz X Y Xy Z Xyz. product(Y::'a,Z,Yz) & product(X::'a,Yz,Xyz) & product(X::'a,Y,Xy) --> product(Xy::'a,Z,Xyz)) &  | 
|
393  | 
(\<forall>Y X Z. defined(X::'a,Y) & defined(Y::'a,Z) & identity_map(Y) --> defined(X::'a,Z)) &  | 
|
394  | 
(\<forall>X. identity_map(domain(X))) &  | 
|
395  | 
(\<forall>X. identity_map(codomain(X))) &  | 
|
396  | 
(\<forall>X. defined(X::'a,domain(X))) &  | 
|
397  | 
(\<forall>X. defined(codomain(X),X)) &  | 
|
398  | 
(\<forall>X. product(X::'a,domain(X),X)) &  | 
|
399  | 
(\<forall>X. product(codomain(X),X,X)) &  | 
|
400  | 
(\<forall>X Y. defined(X::'a,Y) & identity_map(X) --> product(X::'a,Y,Y)) &  | 
|
401  | 
(\<forall>Y X. defined(X::'a,Y) & identity_map(Y) --> product(X::'a,Y,X)) &  | 
|
| 24127 | 402  | 
(\<forall>X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W))"  | 
403  | 
||
404  | 
abbreviation "CAT001_0_eq compos defined identity_map codomain domain product equal \<equiv>  | 
|
| 24128 | 405  | 
(\<forall>X Y Z W. equal(X::'a,Y) & product(X::'a,Z,W) --> product(Y::'a,Z,W)) &  | 
406  | 
(\<forall>X Z Y W. equal(X::'a,Y) & product(Z::'a,X,W) --> product(Z::'a,Y,W)) &  | 
|
407  | 
(\<forall>X Z W Y. equal(X::'a,Y) & product(Z::'a,W,X) --> product(Z::'a,W,Y)) &  | 
|
408  | 
(\<forall>X Y. equal(X::'a,Y) --> equal(domain(X),domain(Y))) &  | 
|
409  | 
(\<forall>X Y. equal(X::'a,Y) --> equal(codomain(X),codomain(Y))) &  | 
|
410  | 
(\<forall>X Y. equal(X::'a,Y) & identity_map(X) --> identity_map(Y)) &  | 
|
411  | 
(\<forall>X Y Z. equal(X::'a,Y) & defined(X::'a,Z) --> defined(Y::'a,Z)) &  | 
|
412  | 
(\<forall>X Z Y. equal(X::'a,Y) & defined(Z::'a,X) --> defined(Z::'a,Y)) &  | 
|
413  | 
(\<forall>X Z Y. equal(X::'a,Y) --> equal(compos(Z::'a,X),compos(Z::'a,Y))) &  | 
|
| 24127 | 414  | 
(\<forall>X Y Z. equal(X::'a,Y) --> equal(compos(X::'a,Z),compos(Y::'a,Z)))"  | 
415  | 
||
416  | 
(*54288 inferences so far. Searching to depth 14. 118.0 secs*)  | 
|
417  | 
lemma CAT005_1:  | 
|
418  | 
"EQU001_0_ax equal &  | 
|
419  | 
CAT001_0_ax equal codomain domain identity_map compos product defined &  | 
|
420  | 
CAT001_0_eq compos defined identity_map codomain domain product equal &  | 
|
| 24128 | 421  | 
(defined(a::'a,d)) &  | 
422  | 
(identity_map(d)) &  | 
|
| 24127 | 423  | 
(~equal(domain(a),d)) --> False"  | 
424  | 
oops  | 
|
425  | 
||
426  | 
(*1728 inferences so far. Searching to depth 10. 5.8 secs*)  | 
|
427  | 
lemma CAT007_1:  | 
|
428  | 
"EQU001_0_ax equal &  | 
|
429  | 
CAT001_0_ax equal codomain domain identity_map compos product defined &  | 
|
430  | 
CAT001_0_eq compos defined identity_map codomain domain product equal &  | 
|
| 24128 | 431  | 
(equal(domain(a),codomain(b))) &  | 
| 24127 | 432  | 
(~defined(a::'a,b)) --> False"  | 
433  | 
by meson  | 
|
434  | 
||
435  | 
(*82895 inferences so far. Searching to depth 13. 355 secs*)  | 
|
436  | 
lemma CAT018_1:  | 
|
437  | 
"EQU001_0_ax equal &  | 
|
438  | 
CAT001_0_ax equal codomain domain identity_map compos product defined &  | 
|
439  | 
CAT001_0_eq compos defined identity_map codomain domain product equal &  | 
|
| 24128 | 440  | 
(defined(a::'a,b)) &  | 
441  | 
(defined(b::'a,c)) &  | 
|
| 24127 | 442  | 
(~defined(a::'a,compos(b::'a,c))) --> False"  | 
443  | 
oops  | 
|
444  | 
||
445  | 
(*1118 inferences so far. Searching to depth 8. 2.3 secs*)  | 
|
446  | 
lemma COL001_2:  | 
|
447  | 
"EQU001_0_ax equal &  | 
|
| 24128 | 448  | 
(\<forall>X Y Z. equal(apply(apply(apply(s::'a,X),Y),Z),apply(apply(X::'a,Z),apply(Y::'a,Z)))) &  | 
449  | 
(\<forall>Y X. equal(apply(apply(k::'a,X),Y),X)) &  | 
|
450  | 
(\<forall>X Y Z. equal(apply(apply(apply(b::'a,X),Y),Z),apply(X::'a,apply(Y::'a,Z)))) &  | 
|
451  | 
(\<forall>X. equal(apply(i::'a,X),X)) &  | 
|
452  | 
(\<forall>A B C. equal(A::'a,B) --> equal(apply(A::'a,C),apply(B::'a,C))) &  | 
|
453  | 
(\<forall>D F' E. equal(D::'a,E) --> equal(apply(F'::'a,D),apply(F'::'a,E))) &  | 
|
454  | 
(\<forall>X. equal(apply(apply(apply(s::'a,apply(b::'a,X)),i),apply(apply(s::'a,apply(b::'a,X)),i)),apply(x::'a,apply(apply(apply(s::'a,apply(b::'a,X)),i),apply(apply(s::'a,apply(b::'a,X)),i))))) &  | 
|
| 24127 | 455  | 
(\<forall>Y. ~equal(Y::'a,apply(combinator::'a,Y))) --> False"  | 
456  | 
by meson  | 
|
457  | 
||
458  | 
(*500 inferences so far. Searching to depth 8. 0.9 secs*)  | 
|
459  | 
lemma COL023_1:  | 
|
460  | 
"EQU001_0_ax equal &  | 
|
| 24128 | 461  | 
(\<forall>X Y Z. equal(apply(apply(apply(b::'a,X),Y),Z),apply(X::'a,apply(Y::'a,Z)))) &  | 
462  | 
(\<forall>X Y Z. equal(apply(apply(apply(n::'a,X),Y),Z),apply(apply(apply(X::'a,Z),Y),Z))) &  | 
|
463  | 
(\<forall>A B C. equal(A::'a,B) --> equal(apply(A::'a,C),apply(B::'a,C))) &  | 
|
464  | 
(\<forall>D F' E. equal(D::'a,E) --> equal(apply(F'::'a,D),apply(F'::'a,E))) &  | 
|
| 24127 | 465  | 
(\<forall>Y. ~equal(Y::'a,apply(combinator::'a,Y))) --> False"  | 
466  | 
by meson  | 
|
467  | 
||
468  | 
(*3018 inferences so far. Searching to depth 10. 4.3 secs*)  | 
|
469  | 
lemma COL032_1:  | 
|
470  | 
"EQU001_0_ax equal &  | 
|
| 24128 | 471  | 
(\<forall>X. equal(apply(m::'a,X),apply(X::'a,X))) &  | 
472  | 
(\<forall>Y X Z. equal(apply(apply(apply(q::'a,X),Y),Z),apply(Y::'a,apply(X::'a,Z)))) &  | 
|
473  | 
(\<forall>A B C. equal(A::'a,B) --> equal(apply(A::'a,C),apply(B::'a,C))) &  | 
|
474  | 
(\<forall>D F' E. equal(D::'a,E) --> equal(apply(F'::'a,D),apply(F'::'a,E))) &  | 
|
475  | 
(\<forall>G H. equal(G::'a,H) --> equal(f(G),f(H))) &  | 
|
| 24127 | 476  | 
(\<forall>Y. ~equal(apply(Y::'a,f(Y)),apply(f(Y),apply(Y::'a,f(Y))))) --> False"  | 
477  | 
by meson  | 
|
478  | 
||
479  | 
(*381878 inferences so far. Searching to depth 13. 670.4 secs*)  | 
|
480  | 
lemma COL052_2:  | 
|
481  | 
"EQU001_0_ax equal &  | 
|
| 24128 | 482  | 
(\<forall>X Y W. equal(response(compos(X::'a,Y),W),response(X::'a,response(Y::'a,W)))) &  | 
483  | 
(\<forall>X Y. agreeable(X) --> equal(response(X::'a,common_bird(Y)),response(Y::'a,common_bird(Y)))) &  | 
|
484  | 
(\<forall>Z X. equal(response(X::'a,Z),response(compatible(X),Z)) --> agreeable(X)) &  | 
|
485  | 
(\<forall>A B. equal(A::'a,B) --> equal(common_bird(A),common_bird(B))) &  | 
|
486  | 
(\<forall>C D. equal(C::'a,D) --> equal(compatible(C),compatible(D))) &  | 
|
487  | 
(\<forall>Q R. equal(Q::'a,R) & agreeable(Q) --> agreeable(R)) &  | 
|
488  | 
(\<forall>A B C. equal(A::'a,B) --> equal(compos(A::'a,C),compos(B::'a,C))) &  | 
|
489  | 
(\<forall>D F' E. equal(D::'a,E) --> equal(compos(F'::'a,D),compos(F'::'a,E))) &  | 
|
490  | 
(\<forall>G H I'. equal(G::'a,H) --> equal(response(G::'a,I'),response(H::'a,I'))) &  | 
|
491  | 
(\<forall>J L K'. equal(J::'a,K') --> equal(response(L::'a,J),response(L::'a,K'))) &  | 
|
492  | 
(agreeable(c)) &  | 
|
493  | 
(~agreeable(a)) &  | 
|
| 24127 | 494  | 
(equal(c::'a,compos(a::'a,b))) --> False"  | 
495  | 
oops  | 
|
496  | 
||
497  | 
(*13201 inferences so far. Searching to depth 11. 31.9 secs*)  | 
|
498  | 
lemma COL075_2:  | 
|
| 24128 | 499  | 
"EQU001_0_ax equal &  | 
500  | 
(\<forall>Y X. equal(apply(apply(k::'a,X),Y),X)) &  | 
|
501  | 
(\<forall>X Y Z. equal(apply(apply(apply(abstraction::'a,X),Y),Z),apply(apply(X::'a,apply(k::'a,Z)),apply(Y::'a,Z)))) &  | 
|
502  | 
(\<forall>D E F'. equal(D::'a,E) --> equal(apply(D::'a,F'),apply(E::'a,F'))) &  | 
|
503  | 
(\<forall>G I' H. equal(G::'a,H) --> equal(apply(I'::'a,G),apply(I'::'a,H))) &  | 
|
504  | 
(\<forall>A B. equal(A::'a,B) --> equal(b(A),b(B))) &  | 
|
505  | 
(\<forall>C D. equal(C::'a,D) --> equal(c(C),c(D))) &  | 
|
| 24127 | 506  | 
(\<forall>Y. ~equal(apply(apply(Y::'a,b(Y)),c(Y)),apply(b(Y),b(Y)))) --> False"  | 
507  | 
oops  | 
|
508  | 
||
509  | 
(*33 inferences so far. Searching to depth 7. 0.1 secs*)  | 
|
510  | 
lemma COM001_1:  | 
|
| 24128 | 511  | 
"(\<forall>Goal_state Start_state. follows(Goal_state::'a,Start_state) --> succeeds(Goal_state::'a,Start_state)) &  | 
512  | 
(\<forall>Goal_state Intermediate_state Start_state. succeeds(Goal_state::'a,Intermediate_state) & succeeds(Intermediate_state::'a,Start_state) --> succeeds(Goal_state::'a,Start_state)) &  | 
|
513  | 
(\<forall>Start_state Label Goal_state. has(Start_state::'a,goto(Label)) & labels(Label::'a,Goal_state) --> succeeds(Goal_state::'a,Start_state)) &  | 
|
514  | 
(\<forall>Start_state Condition Goal_state. has(Start_state::'a,ifthen(Condition::'a,Goal_state)) --> succeeds(Goal_state::'a,Start_state)) &  | 
|
515  | 
(labels(loop::'a,p3)) &  | 
|
516  | 
(has(p3::'a,ifthen(equal(register_j::'a,n),p4))) &  | 
|
517  | 
(has(p4::'a,goto(out))) &  | 
|
518  | 
(follows(p5::'a,p4)) &  | 
|
519  | 
(follows(p8::'a,p3)) &  | 
|
520  | 
(has(p8::'a,goto(loop))) &  | 
|
| 24127 | 521  | 
(~succeeds(p3::'a,p3)) --> False"  | 
522  | 
by meson  | 
|
523  | 
||
524  | 
(*533 inferences so far. Searching to depth 13. 0.3 secs*)  | 
|
525  | 
lemma COM002_1:  | 
|
| 24128 | 526  | 
"(\<forall>Goal_state Start_state. follows(Goal_state::'a,Start_state) --> succeeds(Goal_state::'a,Start_state)) &  | 
527  | 
(\<forall>Goal_state Intermediate_state Start_state. succeeds(Goal_state::'a,Intermediate_state) & succeeds(Intermediate_state::'a,Start_state) --> succeeds(Goal_state::'a,Start_state)) &  | 
|
528  | 
(\<forall>Start_state Label Goal_state. has(Start_state::'a,goto(Label)) & labels(Label::'a,Goal_state) --> succeeds(Goal_state::'a,Start_state)) &  | 
|
529  | 
(\<forall>Start_state Condition Goal_state. has(Start_state::'a,ifthen(Condition::'a,Goal_state)) --> succeeds(Goal_state::'a,Start_state)) &  | 
|
530  | 
(has(p1::'a,assign(register_j::'a,num0))) &  | 
|
531  | 
(follows(p2::'a,p1)) &  | 
|
532  | 
(has(p2::'a,assign(register_k::'a,num1))) &  | 
|
533  | 
(labels(loop::'a,p3)) &  | 
|
534  | 
(follows(p3::'a,p2)) &  | 
|
535  | 
(has(p3::'a,ifthen(equal(register_j::'a,n),p4))) &  | 
|
536  | 
(has(p4::'a,goto(out))) &  | 
|
537  | 
(follows(p5::'a,p4)) &  | 
|
538  | 
(follows(p6::'a,p3)) &  | 
|
539  | 
(has(p6::'a,assign(register_k::'a,mtimes(num2::'a,register_k)))) &  | 
|
540  | 
(follows(p7::'a,p6)) &  | 
|
541  | 
(has(p7::'a,assign(register_j::'a,mplus(register_j::'a,num1)))) &  | 
|
542  | 
(follows(p8::'a,p7)) &  | 
|
543  | 
(has(p8::'a,goto(loop))) &  | 
|
| 24127 | 544  | 
(~succeeds(p3::'a,p3)) --> False"  | 
545  | 
by meson  | 
|
546  | 
||
547  | 
(*4821 inferences so far. Searching to depth 14. 1.3 secs*)  | 
|
548  | 
lemma COM002_2:  | 
|
| 24128 | 549  | 
"(\<forall>Goal_state Start_state. ~(fails(Goal_state::'a,Start_state) & follows(Goal_state::'a,Start_state))) &  | 
550  | 
(\<forall>Goal_state Intermediate_state Start_state. fails(Goal_state::'a,Start_state) --> fails(Goal_state::'a,Intermediate_state) | fails(Intermediate_state::'a,Start_state)) &  | 
|
551  | 
(\<forall>Start_state Label Goal_state. ~(fails(Goal_state::'a,Start_state) & has(Start_state::'a,goto(Label)) & labels(Label::'a,Goal_state))) &  | 
|
552  | 
(\<forall>Start_state Condition Goal_state. ~(fails(Goal_state::'a,Start_state) & has(Start_state::'a,ifthen(Condition::'a,Goal_state)))) &  | 
|
553  | 
(has(p1::'a,assign(register_j::'a,num0))) &  | 
|
554  | 
(follows(p2::'a,p1)) &  | 
|
555  | 
(has(p2::'a,assign(register_k::'a,num1))) &  | 
|
556  | 
(labels(loop::'a,p3)) &  | 
|
557  | 
(follows(p3::'a,p2)) &  | 
|
558  | 
(has(p3::'a,ifthen(equal(register_j::'a,n),p4))) &  | 
|
559  | 
(has(p4::'a,goto(out))) &  | 
|
560  | 
(follows(p5::'a,p4)) &  | 
|
561  | 
(follows(p6::'a,p3)) &  | 
|
562  | 
(has(p6::'a,assign(register_k::'a,mtimes(num2::'a,register_k)))) &  | 
|
563  | 
(follows(p7::'a,p6)) &  | 
|
564  | 
(has(p7::'a,assign(register_j::'a,mplus(register_j::'a,num1)))) &  | 
|
565  | 
(follows(p8::'a,p7)) &  | 
|
566  | 
(has(p8::'a,goto(loop))) &  | 
|
| 24127 | 567  | 
(fails(p3::'a,p3)) --> False"  | 
568  | 
by meson  | 
|
569  | 
||
570  | 
(*98 inferences so far. Searching to depth 10. 1.1 secs*)  | 
|
571  | 
lemma COM003_2:  | 
|
| 24128 | 572  | 
"(\<forall>X Y Z. program_decides(X) & program(Y) --> decides(X::'a,Y,Z)) &  | 
573  | 
(\<forall>X. program_decides(X) | program(f2(X))) &  | 
|
574  | 
(\<forall>X. decides(X::'a,f2(X),f1(X)) --> program_decides(X)) &  | 
|
575  | 
(\<forall>X. program_program_decides(X) --> program(X)) &  | 
|
576  | 
(\<forall>X. program_program_decides(X) --> program_decides(X)) &  | 
|
577  | 
(\<forall>X. program(X) & program_decides(X) --> program_program_decides(X)) &  | 
|
578  | 
(\<forall>X. algorithm_program_decides(X) --> algorithm(X)) &  | 
|
579  | 
(\<forall>X. algorithm_program_decides(X) --> program_decides(X)) &  | 
|
580  | 
(\<forall>X. algorithm(X) & program_decides(X) --> algorithm_program_decides(X)) &  | 
|
581  | 
(\<forall>Y X. program_halts2(X::'a,Y) --> program(X)) &  | 
|
582  | 
(\<forall>X Y. program_halts2(X::'a,Y) --> halts2(X::'a,Y)) &  | 
|
583  | 
(\<forall>X Y. program(X) & halts2(X::'a,Y) --> program_halts2(X::'a,Y)) &  | 
|
584  | 
(\<forall>W X Y Z. halts3_outputs(X::'a,Y,Z,W) --> halts3(X::'a,Y,Z)) &  | 
|
585  | 
(\<forall>Y Z X W. halts3_outputs(X::'a,Y,Z,W) --> outputs(X::'a,W)) &  | 
|
586  | 
(\<forall>Y Z X W. halts3(X::'a,Y,Z) & outputs(X::'a,W) --> halts3_outputs(X::'a,Y,Z,W)) &  | 
|
587  | 
(\<forall>Y X. program_not_halts2(X::'a,Y) --> program(X)) &  | 
|
588  | 
(\<forall>X Y. ~(program_not_halts2(X::'a,Y) & halts2(X::'a,Y))) &  | 
|
589  | 
(\<forall>X Y. program(X) --> program_not_halts2(X::'a,Y) | halts2(X::'a,Y)) &  | 
|
590  | 
(\<forall>W X Y. halts2_outputs(X::'a,Y,W) --> halts2(X::'a,Y)) &  | 
|
591  | 
(\<forall>Y X W. halts2_outputs(X::'a,Y,W) --> outputs(X::'a,W)) &  | 
|
592  | 
(\<forall>Y X W. halts2(X::'a,Y) & outputs(X::'a,W) --> halts2_outputs(X::'a,Y,W)) &  | 
|
593  | 
(\<forall>X W Y Z. program_halts2_halts3_outputs(X::'a,Y,Z,W) --> program_halts2(Y::'a,Z)) &  | 
|
594  | 
(\<forall>X Y Z W. program_halts2_halts3_outputs(X::'a,Y,Z,W) --> halts3_outputs(X::'a,Y,Z,W)) &  | 
|
595  | 
(\<forall>X Y Z W. program_halts2(Y::'a,Z) & halts3_outputs(X::'a,Y,Z,W) --> program_halts2_halts3_outputs(X::'a,Y,Z,W)) &  | 
|
596  | 
(\<forall>X W Y Z. program_not_halts2_halts3_outputs(X::'a,Y,Z,W) --> program_not_halts2(Y::'a,Z)) &  | 
|
597  | 
(\<forall>X Y Z W. program_not_halts2_halts3_outputs(X::'a,Y,Z,W) --> halts3_outputs(X::'a,Y,Z,W)) &  | 
|
598  | 
(\<forall>X Y Z W. program_not_halts2(Y::'a,Z) & halts3_outputs(X::'a,Y,Z,W) --> program_not_halts2_halts3_outputs(X::'a,Y,Z,W)) &  | 
|
599  | 
(\<forall>X W Y. program_halts2_halts2_outputs(X::'a,Y,W) --> program_halts2(Y::'a,Y)) &  | 
|
600  | 
(\<forall>X Y W. program_halts2_halts2_outputs(X::'a,Y,W) --> halts2_outputs(X::'a,Y,W)) &  | 
|
601  | 
(\<forall>X Y W. program_halts2(Y::'a,Y) & halts2_outputs(X::'a,Y,W) --> program_halts2_halts2_outputs(X::'a,Y,W)) &  | 
|
602  | 
(\<forall>X W Y. program_not_halts2_halts2_outputs(X::'a,Y,W) --> program_not_halts2(Y::'a,Y)) &  | 
|
603  | 
(\<forall>X Y W. program_not_halts2_halts2_outputs(X::'a,Y,W) --> halts2_outputs(X::'a,Y,W)) &  | 
|
604  | 
(\<forall>X Y W. program_not_halts2(Y::'a,Y) & halts2_outputs(X::'a,Y,W) --> program_not_halts2_halts2_outputs(X::'a,Y,W)) &  | 
|
605  | 
(\<forall>X. algorithm_program_decides(X) --> program_program_decides(c1)) &  | 
|
606  | 
(\<forall>W Y Z. program_program_decides(W) --> program_halts2_halts3_outputs(W::'a,Y,Z,good)) &  | 
|
607  | 
(\<forall>W Y Z. program_program_decides(W) --> program_not_halts2_halts3_outputs(W::'a,Y,Z,bad)) &  | 
|
608  | 
(\<forall>W. program(W) & program_halts2_halts3_outputs(W::'a,f3(W),f3(W),good) & program_not_halts2_halts3_outputs(W::'a,f3(W),f3(W),bad) --> program(c2)) &  | 
|
609  | 
(\<forall>W Y. program(W) & program_halts2_halts3_outputs(W::'a,f3(W),f3(W),good) & program_not_halts2_halts3_outputs(W::'a,f3(W),f3(W),bad) --> program_halts2_halts2_outputs(c2::'a,Y,good)) &  | 
|
610  | 
(\<forall>W Y. program(W) & program_halts2_halts3_outputs(W::'a,f3(W),f3(W),good) & program_not_halts2_halts3_outputs(W::'a,f3(W),f3(W),bad) --> program_not_halts2_halts2_outputs(c2::'a,Y,bad)) &  | 
|
611  | 
(\<forall>V. program(V) & program_halts2_halts2_outputs(V::'a,f4(V),good) & program_not_halts2_halts2_outputs(V::'a,f4(V),bad) --> program(c3)) &  | 
|
612  | 
(\<forall>V Y. program(V) & program_halts2_halts2_outputs(V::'a,f4(V),good) & program_not_halts2_halts2_outputs(V::'a,f4(V),bad) & program_halts2(Y::'a,Y) --> halts2(c3::'a,Y)) &  | 
|
613  | 
(\<forall>V Y. program(V) & program_halts2_halts2_outputs(V::'a,f4(V),good) & program_not_halts2_halts2_outputs(V::'a,f4(V),bad) --> program_not_halts2_halts2_outputs(c3::'a,Y,bad)) &  | 
|
| 24127 | 614  | 
(algorithm_program_decides(c4)) --> False"  | 
615  | 
by meson  | 
|
616  | 
||
| 24128 | 617  | 
(*2100398 inferences so far. Searching to depth 12.  | 
| 24127 | 618  | 
1256s (21 mins) on griffon*)  | 
619  | 
lemma COM004_1:  | 
|
| 24128 | 620  | 
"EQU001_0_ax equal &  | 
621  | 
(\<forall>C D P Q X Y. failure_node(X::'a,or(C::'a,P)) & failure_node(Y::'a,or(D::'a,Q)) & contradictory(P::'a,Q) & siblings(X::'a,Y) --> failure_node(parent_of(X::'a,Y),or(C::'a,D))) &  | 
|
622  | 
(\<forall>X. contradictory(negate(X),X)) &  | 
|
623  | 
(\<forall>X. contradictory(X::'a,negate(X))) &  | 
|
624  | 
(\<forall>X. siblings(left_child_of(X),right_child_of(X))) &  | 
|
625  | 
(\<forall>D E. equal(D::'a,E) --> equal(left_child_of(D),left_child_of(E))) &  | 
|
626  | 
(\<forall>F' G. equal(F'::'a,G) --> equal(negate(F'),negate(G))) &  | 
|
627  | 
(\<forall>H I' J. equal(H::'a,I') --> equal(or(H::'a,J),or(I'::'a,J))) &  | 
|
628  | 
(\<forall>K' M L. equal(K'::'a,L) --> equal(or(M::'a,K'),or(M::'a,L))) &  | 
|
629  | 
(\<forall>N O' P. equal(N::'a,O') --> equal(parent_of(N::'a,P),parent_of(O'::'a,P))) &  | 
|
630  | 
(\<forall>Q S' R. equal(Q::'a,R) --> equal(parent_of(S'::'a,Q),parent_of(S'::'a,R))) &  | 
|
631  | 
(\<forall>T' U. equal(T'::'a,U) --> equal(right_child_of(T'),right_child_of(U))) &  | 
|
632  | 
(\<forall>V W X. equal(V::'a,W) & contradictory(V::'a,X) --> contradictory(W::'a,X)) &  | 
|
633  | 
(\<forall>Y A1 Z. equal(Y::'a,Z) & contradictory(A1::'a,Y) --> contradictory(A1::'a,Z)) &  | 
|
634  | 
(\<forall>B1 C1 D1. equal(B1::'a,C1) & failure_node(B1::'a,D1) --> failure_node(C1::'a,D1)) &  | 
|
635  | 
(\<forall>E1 G1 F1. equal(E1::'a,F1) & failure_node(G1::'a,E1) --> failure_node(G1::'a,F1)) &  | 
|
636  | 
(\<forall>H1 I1 J1. equal(H1::'a,I1) & siblings(H1::'a,J1) --> siblings(I1::'a,J1)) &  | 
|
637  | 
(\<forall>K1 M1 L1. equal(K1::'a,L1) & siblings(M1::'a,K1) --> siblings(M1::'a,L1)) &  | 
|
638  | 
(failure_node(n_left::'a,or(EMPTY::'a,atom))) &  | 
|
639  | 
(failure_node(n_right::'a,or(EMPTY::'a,negate(atom)))) &  | 
|
640  | 
(equal(n_left::'a,left_child_of(n))) &  | 
|
641  | 
(equal(n_right::'a,right_child_of(n))) &  | 
|
| 24127 | 642  | 
(\<forall>Z. ~failure_node(Z::'a,or(EMPTY::'a,EMPTY))) --> False"  | 
643  | 
oops  | 
|
644  | 
||
645  | 
abbreviation "GEO001_0_ax continuous lower_dimension_point_3 lower_dimension_point_2  | 
|
646  | 
lower_dimension_point_1 extension euclid2 euclid1 outer_pasch equidistant equal between \<equiv>  | 
|
| 24128 | 647  | 
(\<forall>X Y. between(X::'a,Y,X) --> equal(X::'a,Y)) &  | 
648  | 
(\<forall>V X Y Z. between(X::'a,Y,V) & between(Y::'a,Z,V) --> between(X::'a,Y,Z)) &  | 
|
649  | 
(\<forall>Y X V Z. between(X::'a,Y,Z) & between(X::'a,Y,V) --> equal(X::'a,Y) | between(X::'a,Z,V) | between(X::'a,V,Z)) &  | 
|
650  | 
(\<forall>Y X. equidistant(X::'a,Y,Y,X)) &  | 
|
651  | 
(\<forall>Z X Y. equidistant(X::'a,Y,Z,Z) --> equal(X::'a,Y)) &  | 
|
652  | 
(\<forall>X Y Z V V2 W. equidistant(X::'a,Y,Z,V) & equidistant(X::'a,Y,V2,W) --> equidistant(Z::'a,V,V2,W)) &  | 
|
653  | 
(\<forall>W X Z V Y. between(X::'a,W,V) & between(Y::'a,V,Z) --> between(X::'a,outer_pasch(W::'a,X,Y,Z,V),Y)) &  | 
|
654  | 
(\<forall>W X Y Z V. between(X::'a,W,V) & between(Y::'a,V,Z) --> between(Z::'a,W,outer_pasch(W::'a,X,Y,Z,V))) &  | 
|
655  | 
(\<forall>W X Y Z V. between(X::'a,V,W) & between(Y::'a,V,Z) --> equal(X::'a,V) | between(X::'a,Z,euclid1(W::'a,X,Y,Z,V))) &  | 
|
656  | 
(\<forall>W X Y Z V. between(X::'a,V,W) & between(Y::'a,V,Z) --> equal(X::'a,V) | between(X::'a,Y,euclid2(W::'a,X,Y,Z,V))) &  | 
|
657  | 
(\<forall>W X Y Z V. between(X::'a,V,W) & between(Y::'a,V,Z) --> equal(X::'a,V) | between(euclid1(W::'a,X,Y,Z,V),W,euclid2(W::'a,X,Y,Z,V))) &  | 
|
658  | 
(\<forall>X1 Y1 X Y Z V Z1 V1. equidistant(X::'a,Y,X1,Y1) & equidistant(Y::'a,Z,Y1,Z1) & equidistant(X::'a,V,X1,V1) & equidistant(Y::'a,V,Y1,V1) & between(X::'a,Y,Z) & between(X1::'a,Y1,Z1) --> equal(X::'a,Y) | equidistant(Z::'a,V,Z1,V1)) &  | 
|
659  | 
(\<forall>X Y W V. between(X::'a,Y,extension(X::'a,Y,W,V))) &  | 
|
660  | 
(\<forall>X Y W V. equidistant(Y::'a,extension(X::'a,Y,W,V),W,V)) &  | 
|
661  | 
(~between(lower_dimension_point_1::'a,lower_dimension_point_2,lower_dimension_point_3)) &  | 
|
662  | 
(~between(lower_dimension_point_2::'a,lower_dimension_point_3,lower_dimension_point_1)) &  | 
|
663  | 
(~between(lower_dimension_point_3::'a,lower_dimension_point_1,lower_dimension_point_2)) &  | 
|
664  | 
(\<forall>Z X Y W V. equidistant(X::'a,W,X,V) & equidistant(Y::'a,W,Y,V) & equidistant(Z::'a,W,Z,V) --> between(X::'a,Y,Z) | between(Y::'a,Z,X) | between(Z::'a,X,Y) | equal(W::'a,V)) &  | 
|
665  | 
(\<forall>X Y Z X1 Z1 V. equidistant(V::'a,X,V,X1) & equidistant(V::'a,Z,V,Z1) & between(V::'a,X,Z) & between(X::'a,Y,Z) --> equidistant(V::'a,Y,Z,continuous(X::'a,Y,Z,X1,Z1,V))) &  | 
|
| 24127 | 666  | 
(\<forall>X Y Z X1 V Z1. equidistant(V::'a,X,V,X1) & equidistant(V::'a,Z,V,Z1) & between(V::'a,X,Z) & between(X::'a,Y,Z) --> between(X1::'a,continuous(X::'a,Y,Z,X1,Z1,V),Z1))"  | 
667  | 
||
668  | 
abbreviation "GEO001_0_eq continuous extension euclid2 euclid1 outer_pasch equidistant  | 
|
669  | 
between equal \<equiv>  | 
|
| 24128 | 670  | 
(\<forall>X Y W Z. equal(X::'a,Y) & between(X::'a,W,Z) --> between(Y::'a,W,Z)) &  | 
671  | 
(\<forall>X W Y Z. equal(X::'a,Y) & between(W::'a,X,Z) --> between(W::'a,Y,Z)) &  | 
|
672  | 
(\<forall>X W Z Y. equal(X::'a,Y) & between(W::'a,Z,X) --> between(W::'a,Z,Y)) &  | 
|
673  | 
(\<forall>X Y V W Z. equal(X::'a,Y) & equidistant(X::'a,V,W,Z) --> equidistant(Y::'a,V,W,Z)) &  | 
|
674  | 
(\<forall>X V Y W Z. equal(X::'a,Y) & equidistant(V::'a,X,W,Z) --> equidistant(V::'a,Y,W,Z)) &  | 
|
675  | 
(\<forall>X V W Y Z. equal(X::'a,Y) & equidistant(V::'a,W,X,Z) --> equidistant(V::'a,W,Y,Z)) &  | 
|
676  | 
(\<forall>X V W Z Y. equal(X::'a,Y) & equidistant(V::'a,W,Z,X) --> equidistant(V::'a,W,Z,Y)) &  | 
|
677  | 
(\<forall>X Y V1 V2 V3 V4. equal(X::'a,Y) --> equal(outer_pasch(X::'a,V1,V2,V3,V4),outer_pasch(Y::'a,V1,V2,V3,V4))) &  | 
|
678  | 
(\<forall>X V1 Y V2 V3 V4. equal(X::'a,Y) --> equal(outer_pasch(V1::'a,X,V2,V3,V4),outer_pasch(V1::'a,Y,V2,V3,V4))) &  | 
|
679  | 
(\<forall>X V1 V2 Y V3 V4. equal(X::'a,Y) --> equal(outer_pasch(V1::'a,V2,X,V3,V4),outer_pasch(V1::'a,V2,Y,V3,V4))) &  | 
|
680  | 
(\<forall>X V1 V2 V3 Y V4. equal(X::'a,Y) --> equal(outer_pasch(V1::'a,V2,V3,X,V4),outer_pasch(V1::'a,V2,V3,Y,V4))) &  | 
|
681  | 
(\<forall>X V1 V2 V3 V4 Y. equal(X::'a,Y) --> equal(outer_pasch(V1::'a,V2,V3,V4,X),outer_pasch(V1::'a,V2,V3,V4,Y))) &  | 
|
682  | 
(\<forall>A B C D E F'. equal(A::'a,B) --> equal(euclid1(A::'a,C,D,E,F'),euclid1(B::'a,C,D,E,F'))) &  | 
|
683  | 
(\<forall>G I' H J K' L. equal(G::'a,H) --> equal(euclid1(I'::'a,G,J,K',L),euclid1(I'::'a,H,J,K',L))) &  | 
|
684  | 
(\<forall>M O' P N Q R. equal(M::'a,N) --> equal(euclid1(O'::'a,P,M,Q,R),euclid1(O'::'a,P,N,Q,R))) &  | 
|
685  | 
(\<forall>S' U V W T' X. equal(S'::'a,T') --> equal(euclid1(U::'a,V,W,S',X),euclid1(U::'a,V,W,T',X))) &  | 
|
686  | 
(\<forall>Y A1 B1 C1 D1 Z. equal(Y::'a,Z) --> equal(euclid1(A1::'a,B1,C1,D1,Y),euclid1(A1::'a,B1,C1,D1,Z))) &  | 
|
687  | 
(\<forall>E1 F1 G1 H1 I1 J1. equal(E1::'a,F1) --> equal(euclid2(E1::'a,G1,H1,I1,J1),euclid2(F1::'a,G1,H1,I1,J1))) &  | 
|
688  | 
(\<forall>K1 M1 L1 N1 O1 P1. equal(K1::'a,L1) --> equal(euclid2(M1::'a,K1,N1,O1,P1),euclid2(M1::'a,L1,N1,O1,P1))) &  | 
|
689  | 
(\<forall>Q1 S1 T1 R1 U1 V1. equal(Q1::'a,R1) --> equal(euclid2(S1::'a,T1,Q1,U1,V1),euclid2(S1::'a,T1,R1,U1,V1))) &  | 
|
690  | 
(\<forall>W1 Y1 Z1 A2 X1 B2. equal(W1::'a,X1) --> equal(euclid2(Y1::'a,Z1,A2,W1,B2),euclid2(Y1::'a,Z1,A2,X1,B2))) &  | 
|
691  | 
(\<forall>C2 E2 F2 G2 H2 D2. equal(C2::'a,D2) --> equal(euclid2(E2::'a,F2,G2,H2,C2),euclid2(E2::'a,F2,G2,H2,D2))) &  | 
|
692  | 
(\<forall>X Y V1 V2 V3. equal(X::'a,Y) --> equal(extension(X::'a,V1,V2,V3),extension(Y::'a,V1,V2,V3))) &  | 
|
693  | 
(\<forall>X V1 Y V2 V3. equal(X::'a,Y) --> equal(extension(V1::'a,X,V2,V3),extension(V1::'a,Y,V2,V3))) &  | 
|
694  | 
(\<forall>X V1 V2 Y V3. equal(X::'a,Y) --> equal(extension(V1::'a,V2,X,V3),extension(V1::'a,V2,Y,V3))) &  | 
|
695  | 
(\<forall>X V1 V2 V3 Y. equal(X::'a,Y) --> equal(extension(V1::'a,V2,V3,X),extension(V1::'a,V2,V3,Y))) &  | 
|
696  | 
(\<forall>X Y V1 V2 V3 V4 V5. equal(X::'a,Y) --> equal(continuous(X::'a,V1,V2,V3,V4,V5),continuous(Y::'a,V1,V2,V3,V4,V5))) &  | 
|
697  | 
(\<forall>X V1 Y V2 V3 V4 V5. equal(X::'a,Y) --> equal(continuous(V1::'a,X,V2,V3,V4,V5),continuous(V1::'a,Y,V2,V3,V4,V5))) &  | 
|
698  | 
(\<forall>X V1 V2 Y V3 V4 V5. equal(X::'a,Y) --> equal(continuous(V1::'a,V2,X,V3,V4,V5),continuous(V1::'a,V2,Y,V3,V4,V5))) &  | 
|
699  | 
(\<forall>X V1 V2 V3 Y V4 V5. equal(X::'a,Y) --> equal(continuous(V1::'a,V2,V3,X,V4,V5),continuous(V1::'a,V2,V3,Y,V4,V5))) &  | 
|
700  | 
(\<forall>X V1 V2 V3 V4 Y V5. equal(X::'a,Y) --> equal(continuous(V1::'a,V2,V3,V4,X,V5),continuous(V1::'a,V2,V3,V4,Y,V5))) &  | 
|
| 24127 | 701  | 
(\<forall>X V1 V2 V3 V4 V5 Y. equal(X::'a,Y) --> equal(continuous(V1::'a,V2,V3,V4,V5,X),continuous(V1::'a,V2,V3,V4,V5,Y)))"  | 
702  | 
||
703  | 
||
704  | 
(*179 inferences so far. Searching to depth 7. 3.9 secs*)  | 
|
705  | 
lemma GEO003_1:  | 
|
706  | 
"EQU001_0_ax equal &  | 
|
707  | 
GEO001_0_ax continuous lower_dimension_point_3 lower_dimension_point_2  | 
|
708  | 
lower_dimension_point_1 extension euclid2 euclid1 outer_pasch equidistant equal between &  | 
|
709  | 
GEO001_0_eq continuous extension euclid2 euclid1 outer_pasch equidistant between equal &  | 
|
710  | 
(~between(a::'a,b,b)) --> False"  | 
|
711  | 
by meson  | 
|
712  | 
||
713  | 
abbreviation "GEO002_ax_eq continuous euclid2 euclid1 lower_dimension_point_3  | 
|
714  | 
lower_dimension_point_2 lower_dimension_point_1 inner_pasch extension  | 
|
715  | 
between equal equidistant \<equiv>  | 
|
| 24128 | 716  | 
(\<forall>Y X. equidistant(X::'a,Y,Y,X)) &  | 
717  | 
(\<forall>X Y Z V V2 W. equidistant(X::'a,Y,Z,V) & equidistant(X::'a,Y,V2,W) --> equidistant(Z::'a,V,V2,W)) &  | 
|
718  | 
(\<forall>Z X Y. equidistant(X::'a,Y,Z,Z) --> equal(X::'a,Y)) &  | 
|
719  | 
(\<forall>X Y W V. between(X::'a,Y,extension(X::'a,Y,W,V))) &  | 
|
720  | 
(\<forall>X Y W V. equidistant(Y::'a,extension(X::'a,Y,W,V),W,V)) &  | 
|
721  | 
(\<forall>X1 Y1 X Y Z V Z1 V1. equidistant(X::'a,Y,X1,Y1) & equidistant(Y::'a,Z,Y1,Z1) & equidistant(X::'a,V,X1,V1) & equidistant(Y::'a,V,Y1,V1) & between(X::'a,Y,Z) & between(X1::'a,Y1,Z1) --> equal(X::'a,Y) | equidistant(Z::'a,V,Z1,V1)) &  | 
|
722  | 
(\<forall>X Y. between(X::'a,Y,X) --> equal(X::'a,Y)) &  | 
|
723  | 
(\<forall>U V W X Y. between(U::'a,V,W) & between(Y::'a,X,W) --> between(V::'a,inner_pasch(U::'a,V,W,X,Y),Y)) &  | 
|
724  | 
(\<forall>V W X Y U. between(U::'a,V,W) & between(Y::'a,X,W) --> between(X::'a,inner_pasch(U::'a,V,W,X,Y),U)) &  | 
|
725  | 
(~between(lower_dimension_point_1::'a,lower_dimension_point_2,lower_dimension_point_3)) &  | 
|
726  | 
(~between(lower_dimension_point_2::'a,lower_dimension_point_3,lower_dimension_point_1)) &  | 
|
727  | 
(~between(lower_dimension_point_3::'a,lower_dimension_point_1,lower_dimension_point_2)) &  | 
|
728  | 
(\<forall>Z X Y W V. equidistant(X::'a,W,X,V) & equidistant(Y::'a,W,Y,V) & equidistant(Z::'a,W,Z,V) --> between(X::'a,Y,Z) | between(Y::'a,Z,X) | between(Z::'a,X,Y) | equal(W::'a,V)) &  | 
|
729  | 
(\<forall>U V W X Y. between(U::'a,W,Y) & between(V::'a,W,X) --> equal(U::'a,W) | between(U::'a,V,euclid1(U::'a,V,W,X,Y))) &  | 
|
730  | 
(\<forall>U V W X Y. between(U::'a,W,Y) & between(V::'a,W,X) --> equal(U::'a,W) | between(U::'a,X,euclid2(U::'a,V,W,X,Y))) &  | 
|
731  | 
(\<forall>U V W X Y. between(U::'a,W,Y) & between(V::'a,W,X) --> equal(U::'a,W) | between(euclid1(U::'a,V,W,X,Y),Y,euclid2(U::'a,V,W,X,Y))) &  | 
|
732  | 
(\<forall>U V V1 W X X1. equidistant(U::'a,V,U,V1) & equidistant(U::'a,X,U,X1) & between(U::'a,V,X) & between(V::'a,W,X) --> between(V1::'a,continuous(U::'a,V,V1,W,X,X1),X1)) &  | 
|
733  | 
(\<forall>U V V1 W X X1. equidistant(U::'a,V,U,V1) & equidistant(U::'a,X,U,X1) & between(U::'a,V,X) & between(V::'a,W,X) --> equidistant(U::'a,W,U,continuous(U::'a,V,V1,W,X,X1))) &  | 
|
734  | 
(\<forall>X Y W Z. equal(X::'a,Y) & between(X::'a,W,Z) --> between(Y::'a,W,Z)) &  | 
|
735  | 
(\<forall>X W Y Z. equal(X::'a,Y) & between(W::'a,X,Z) --> between(W::'a,Y,Z)) &  | 
|
736  | 
(\<forall>X W Z Y. equal(X::'a,Y) & between(W::'a,Z,X) --> between(W::'a,Z,Y)) &  | 
|
737  | 
(\<forall>X Y V W Z. equal(X::'a,Y) & equidistant(X::'a,V,W,Z) --> equidistant(Y::'a,V,W,Z)) &  | 
|
738  | 
(\<forall>X V Y W Z. equal(X::'a,Y) & equidistant(V::'a,X,W,Z) --> equidistant(V::'a,Y,W,Z)) &  | 
|
739  | 
(\<forall>X V W Y Z. equal(X::'a,Y) & equidistant(V::'a,W,X,Z) --> equidistant(V::'a,W,Y,Z)) &  | 
|
740  | 
(\<forall>X V W Z Y. equal(X::'a,Y) & equidistant(V::'a,W,Z,X) --> equidistant(V::'a,W,Z,Y)) &  | 
|
741  | 
(\<forall>X Y V1 V2 V3 V4. equal(X::'a,Y) --> equal(inner_pasch(X::'a,V1,V2,V3,V4),inner_pasch(Y::'a,V1,V2,V3,V4))) &  | 
|
742  | 
(\<forall>X V1 Y V2 V3 V4. equal(X::'a,Y) --> equal(inner_pasch(V1::'a,X,V2,V3,V4),inner_pasch(V1::'a,Y,V2,V3,V4))) &  | 
|
743  | 
(\<forall>X V1 V2 Y V3 V4. equal(X::'a,Y) --> equal(inner_pasch(V1::'a,V2,X,V3,V4),inner_pasch(V1::'a,V2,Y,V3,V4))) &  | 
|
744  | 
(\<forall>X V1 V2 V3 Y V4. equal(X::'a,Y) --> equal(inner_pasch(V1::'a,V2,V3,X,V4),inner_pasch(V1::'a,V2,V3,Y,V4))) &  | 
|
745  | 
(\<forall>X V1 V2 V3 V4 Y. equal(X::'a,Y) --> equal(inner_pasch(V1::'a,V2,V3,V4,X),inner_pasch(V1::'a,V2,V3,V4,Y))) &  | 
|
746  | 
(\<forall>A B C D E F'. equal(A::'a,B) --> equal(euclid1(A::'a,C,D,E,F'),euclid1(B::'a,C,D,E,F'))) &  | 
|
747  | 
(\<forall>G I' H J K' L. equal(G::'a,H) --> equal(euclid1(I'::'a,G,J,K',L),euclid1(I'::'a,H,J,K',L))) &  | 
|
748  | 
(\<forall>M O' P N Q R. equal(M::'a,N) --> equal(euclid1(O'::'a,P,M,Q,R),euclid1(O'::'a,P,N,Q,R))) &  | 
|
749  | 
(\<forall>S' U V W T' X. equal(S'::'a,T') --> equal(euclid1(U::'a,V,W,S',X),euclid1(U::'a,V,W,T',X))) &  | 
|
750  | 
(\<forall>Y A1 B1 C1 D1 Z. equal(Y::'a,Z) --> equal(euclid1(A1::'a,B1,C1,D1,Y),euclid1(A1::'a,B1,C1,D1,Z))) &  | 
|
751  | 
(\<forall>E1 F1 G1 H1 I1 J1. equal(E1::'a,F1) --> equal(euclid2(E1::'a,G1,H1,I1,J1),euclid2(F1::'a,G1,H1,I1,J1))) &  | 
|
752  | 
(\<forall>K1 M1 L1 N1 O1 P1. equal(K1::'a,L1) --> equal(euclid2(M1::'a,K1,N1,O1,P1),euclid2(M1::'a,L1,N1,O1,P1))) &  | 
|
753  | 
(\<forall>Q1 S1 T1 R1 U1 V1. equal(Q1::'a,R1) --> equal(euclid2(S1::'a,T1,Q1,U1,V1),euclid2(S1::'a,T1,R1,U1,V1))) &  | 
|
754  | 
(\<forall>W1 Y1 Z1 A2 X1 B2. equal(W1::'a,X1) --> equal(euclid2(Y1::'a,Z1,A2,W1,B2),euclid2(Y1::'a,Z1,A2,X1,B2))) &  | 
|
755  | 
(\<forall>C2 E2 F2 G2 H2 D2. equal(C2::'a,D2) --> equal(euclid2(E2::'a,F2,G2,H2,C2),euclid2(E2::'a,F2,G2,H2,D2))) &  | 
|
756  | 
(\<forall>X Y V1 V2 V3. equal(X::'a,Y) --> equal(extension(X::'a,V1,V2,V3),extension(Y::'a,V1,V2,V3))) &  | 
|
757  | 
(\<forall>X V1 Y V2 V3. equal(X::'a,Y) --> equal(extension(V1::'a,X,V2,V3),extension(V1::'a,Y,V2,V3))) &  | 
|
758  | 
(\<forall>X V1 V2 Y V3. equal(X::'a,Y) --> equal(extension(V1::'a,V2,X,V3),extension(V1::'a,V2,Y,V3))) &  | 
|
759  | 
(\<forall>X V1 V2 V3 Y. equal(X::'a,Y) --> equal(extension(V1::'a,V2,V3,X),extension(V1::'a,V2,V3,Y))) &  | 
|
760  | 
(\<forall>X Y V1 V2 V3 V4 V5. equal(X::'a,Y) --> equal(continuous(X::'a,V1,V2,V3,V4,V5),continuous(Y::'a,V1,V2,V3,V4,V5))) &  | 
|
761  | 
(\<forall>X V1 Y V2 V3 V4 V5. equal(X::'a,Y) --> equal(continuous(V1::'a,X,V2,V3,V4,V5),continuous(V1::'a,Y,V2,V3,V4,V5))) &  | 
|
762  | 
(\<forall>X V1 V2 Y V3 V4 V5. equal(X::'a,Y) --> equal(continuous(V1::'a,V2,X,V3,V4,V5),continuous(V1::'a,V2,Y,V3,V4,V5))) &  | 
|
763  | 
(\<forall>X V1 V2 V3 Y V4 V5. equal(X::'a,Y) --> equal(continuous(V1::'a,V2,V3,X,V4,V5),continuous(V1::'a,V2,V3,Y,V4,V5))) &  | 
|
764  | 
(\<forall>X V1 V2 V3 V4 Y V5. equal(X::'a,Y) --> equal(continuous(V1::'a,V2,V3,V4,X,V5),continuous(V1::'a,V2,V3,V4,Y,V5))) &  | 
|
| 24127 | 765  | 
(\<forall>X V1 V2 V3 V4 V5 Y. equal(X::'a,Y) --> equal(continuous(V1::'a,V2,V3,V4,V5,X),continuous(V1::'a,V2,V3,V4,V5,Y)))"  | 
766  | 
||
767  | 
(*4272 inferences so far. Searching to depth 10. 29.4 secs*)  | 
|
768  | 
lemma GEO017_2:  | 
|
769  | 
"EQU001_0_ax equal &  | 
|
770  | 
GEO002_ax_eq continuous euclid2 euclid1 lower_dimension_point_3  | 
|
771  | 
lower_dimension_point_2 lower_dimension_point_1 inner_pasch extension  | 
|
772  | 
between equal equidistant &  | 
|
| 24128 | 773  | 
(equidistant(u::'a,v,w,x)) &  | 
| 24127 | 774  | 
(~equidistant(u::'a,v,x,w)) --> False"  | 
775  | 
oops  | 
|
776  | 
||
777  | 
(*382903 inferences so far. Searching to depth 9. 1022s (17 mins) on griffon*)  | 
|
778  | 
lemma GEO027_3:  | 
|
779  | 
"EQU001_0_ax equal &  | 
|
780  | 
GEO002_ax_eq continuous euclid2 euclid1 lower_dimension_point_3  | 
|
781  | 
lower_dimension_point_2 lower_dimension_point_1 inner_pasch extension  | 
|
782  | 
between equal equidistant &  | 
|
| 24128 | 783  | 
(\<forall>U V. equal(reflection(U::'a,V),extension(U::'a,V,U,V))) &  | 
784  | 
(\<forall>X Y Z. equal(X::'a,Y) --> equal(reflection(X::'a,Z),reflection(Y::'a,Z))) &  | 
|
785  | 
(\<forall>A1 C1 B1. equal(A1::'a,B1) --> equal(reflection(C1::'a,A1),reflection(C1::'a,B1))) &  | 
|
786  | 
(\<forall>U V. equidistant(U::'a,V,U,V)) &  | 
|
787  | 
(\<forall>W X U V. equidistant(U::'a,V,W,X) --> equidistant(W::'a,X,U,V)) &  | 
|
788  | 
(\<forall>V U W X. equidistant(U::'a,V,W,X) --> equidistant(V::'a,U,W,X)) &  | 
|
789  | 
(\<forall>U V X W. equidistant(U::'a,V,W,X) --> equidistant(U::'a,V,X,W)) &  | 
|
790  | 
(\<forall>V U X W. equidistant(U::'a,V,W,X) --> equidistant(V::'a,U,X,W)) &  | 
|
791  | 
(\<forall>W X V U. equidistant(U::'a,V,W,X) --> equidistant(W::'a,X,V,U)) &  | 
|
792  | 
(\<forall>X W U V. equidistant(U::'a,V,W,X) --> equidistant(X::'a,W,U,V)) &  | 
|
793  | 
(\<forall>X W V U. equidistant(U::'a,V,W,X) --> equidistant(X::'a,W,V,U)) &  | 
|
794  | 
(\<forall>W X U V Y Z. equidistant(U::'a,V,W,X) & equidistant(W::'a,X,Y,Z) --> equidistant(U::'a,V,Y,Z)) &  | 
|
795  | 
(\<forall>U V W. equal(V::'a,extension(U::'a,V,W,W))) &  | 
|
796  | 
(\<forall>W X U V Y. equal(Y::'a,extension(U::'a,V,W,X)) --> between(U::'a,V,Y)) &  | 
|
797  | 
(\<forall>U V. between(U::'a,V,reflection(U::'a,V))) &  | 
|
798  | 
(\<forall>U V. equidistant(V::'a,reflection(U::'a,V),U,V)) &  | 
|
799  | 
(\<forall>U V. equal(U::'a,V) --> equal(V::'a,reflection(U::'a,V))) &  | 
|
800  | 
(\<forall>U. equal(U::'a,reflection(U::'a,U))) &  | 
|
801  | 
(\<forall>U V. equal(V::'a,reflection(U::'a,V)) --> equal(U::'a,V)) &  | 
|
802  | 
(\<forall>U V. equidistant(U::'a,U,V,V)) &  | 
|
803  | 
(\<forall>V V1 U W U1 W1. equidistant(U::'a,V,U1,V1) & equidistant(V::'a,W,V1,W1) & between(U::'a,V,W) & between(U1::'a,V1,W1) --> equidistant(U::'a,W,U1,W1)) &  | 
|
804  | 
(\<forall>U V W X. between(U::'a,V,W) & between(U::'a,V,X) & equidistant(V::'a,W,V,X) --> equal(U::'a,V) | equal(W::'a,X)) &  | 
|
805  | 
(between(u::'a,v,w)) &  | 
|
806  | 
(~equal(u::'a,v)) &  | 
|
| 24127 | 807  | 
(~equal(w::'a,extension(u::'a,v,v,w))) --> False"  | 
808  | 
oops  | 
|
809  | 
||
810  | 
(*313884 inferences so far. Searching to depth 10. 887 secs: 15 mins.*)  | 
|
811  | 
lemma GEO058_2:  | 
|
812  | 
"EQU001_0_ax equal &  | 
|
813  | 
GEO002_ax_eq continuous euclid2 euclid1 lower_dimension_point_3  | 
|
814  | 
lower_dimension_point_2 lower_dimension_point_1 inner_pasch extension  | 
|
815  | 
between equal equidistant &  | 
|
| 24128 | 816  | 
(\<forall>U V. equal(reflection(U::'a,V),extension(U::'a,V,U,V))) &  | 
817  | 
(\<forall>X Y Z. equal(X::'a,Y) --> equal(reflection(X::'a,Z),reflection(Y::'a,Z))) &  | 
|
818  | 
(\<forall>A1 C1 B1. equal(A1::'a,B1) --> equal(reflection(C1::'a,A1),reflection(C1::'a,B1))) &  | 
|
819  | 
(equal(v::'a,reflection(u::'a,v))) &  | 
|
| 24127 | 820  | 
(~equal(u::'a,v)) --> False"  | 
821  | 
oops  | 
|
822  | 
||
823  | 
(*0 inferences so far. Searching to depth 0. 0.2 secs*)  | 
|
824  | 
lemma GEO079_1:  | 
|
| 24128 | 825  | 
"(\<forall>U V W X Y Z. right_angle(U::'a,V,W) & right_angle(X::'a,Y,Z) --> eq(U::'a,V,W,X,Y,Z)) &  | 
826  | 
(\<forall>U V W X Y Z. CONGRUENT(U::'a,V,W,X,Y,Z) --> eq(U::'a,V,W,X,Y,Z)) &  | 
|
827  | 
(\<forall>V W U X. trapezoid(U::'a,V,W,X) --> parallel(V::'a,W,U,X)) &  | 
|
828  | 
(\<forall>U V X Y. parallel(U::'a,V,X,Y) --> eq(X::'a,V,U,V,X,Y)) &  | 
|
829  | 
(trapezoid(a::'a,b,c,d)) &  | 
|
| 24127 | 830  | 
(~eq(a::'a,c,b,c,a,d)) --> False"  | 
831  | 
by meson  | 
|
832  | 
||
833  | 
abbreviation "GRP003_0_ax equal multiply INVERSE identity product \<equiv>  | 
|
| 24128 | 834  | 
(\<forall>X. product(identity::'a,X,X)) &  | 
835  | 
(\<forall>X. product(X::'a,identity,X)) &  | 
|
836  | 
(\<forall>X. product(INVERSE(X),X,identity)) &  | 
|
837  | 
(\<forall>X. product(X::'a,INVERSE(X),identity)) &  | 
|
838  | 
(\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &  | 
|
839  | 
(\<forall>X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W)) &  | 
|
840  | 
(\<forall>Y U Z X V W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(U::'a,Z,W) --> product(X::'a,V,W)) &  | 
|
| 24127 | 841  | 
(\<forall>Y X V U Z W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(X::'a,V,W) --> product(U::'a,Z,W))"  | 
842  | 
||
843  | 
abbreviation "GRP003_0_eq product multiply INVERSE equal \<equiv>  | 
|
| 24128 | 844  | 
(\<forall>X Y. equal(X::'a,Y) --> equal(INVERSE(X),INVERSE(Y))) &  | 
845  | 
(\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &  | 
|
846  | 
(\<forall>X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) &  | 
|
847  | 
(\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &  | 
|
848  | 
(\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &  | 
|
| 24127 | 849  | 
(\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y))"  | 
850  | 
||
851  | 
(*2032008 inferences so far. Searching to depth 16. 658s (11 mins) on griffon*)  | 
|
852  | 
lemma GRP001_1:  | 
|
853  | 
"EQU001_0_ax equal &  | 
|
854  | 
GRP003_0_ax equal multiply INVERSE identity product &  | 
|
855  | 
GRP003_0_eq product multiply INVERSE equal &  | 
|
| 24128 | 856  | 
(\<forall>X. product(X::'a,X,identity)) &  | 
857  | 
(product(a::'a,b,c)) &  | 
|
| 24127 | 858  | 
(~product(b::'a,a,c)) --> False"  | 
859  | 
oops  | 
|
860  | 
||
861  | 
(*2386 inferences so far. Searching to depth 11. 8.7 secs*)  | 
|
862  | 
lemma GRP008_1:  | 
|
863  | 
"EQU001_0_ax equal &  | 
|
864  | 
GRP003_0_ax equal multiply INVERSE identity product &  | 
|
865  | 
GRP003_0_eq product multiply INVERSE equal &  | 
|
| 24128 | 866  | 
(\<forall>A B. equal(A::'a,B) --> equal(h(A),h(B))) &  | 
867  | 
(\<forall>C D. equal(C::'a,D) --> equal(j(C),j(D))) &  | 
|
868  | 
(\<forall>A B. equal(A::'a,B) & q(A) --> q(B)) &  | 
|
869  | 
(\<forall>B A C. q(A) & product(A::'a,B,C) --> product(B::'a,A,C)) &  | 
|
870  | 
(\<forall>A. product(j(A),A,h(A)) | product(A::'a,j(A),h(A)) | q(A)) &  | 
|
871  | 
(\<forall>A. product(j(A),A,h(A)) & product(A::'a,j(A),h(A)) --> q(A)) &  | 
|
| 24127 | 872  | 
(~q(identity)) --> False"  | 
873  | 
by meson  | 
|
874  | 
||
875  | 
(*8625 inferences so far. Searching to depth 11. 20 secs*)  | 
|
876  | 
lemma GRP013_1:  | 
|
877  | 
"EQU001_0_ax equal &  | 
|
878  | 
GRP003_0_ax equal multiply INVERSE identity product &  | 
|
879  | 
GRP003_0_eq product multiply INVERSE equal &  | 
|
| 24128 | 880  | 
(\<forall>A. product(A::'a,A,identity)) &  | 
881  | 
(product(a::'a,b,c)) &  | 
|
882  | 
(product(INVERSE(a),INVERSE(b),d)) &  | 
|
883  | 
(\<forall>A C B. product(INVERSE(A),INVERSE(B),C) --> product(A::'a,C,B)) &  | 
|
| 24127 | 884  | 
(~product(c::'a,d,identity)) --> False"  | 
885  | 
oops  | 
|
886  | 
||
887  | 
(*2448 inferences so far. Searching to depth 10. 7.2 secs*)  | 
|
888  | 
lemma GRP037_3:  | 
|
889  | 
"EQU001_0_ax equal &  | 
|
890  | 
GRP003_0_ax equal multiply INVERSE identity product &  | 
|
891  | 
GRP003_0_eq product multiply INVERSE equal &  | 
|
| 24128 | 892  | 
(\<forall>A B C. subgroup_member(A) & subgroup_member(B) & product(A::'a,INVERSE(B),C) --> subgroup_member(C)) &  | 
893  | 
(\<forall>A B. equal(A::'a,B) & subgroup_member(A) --> subgroup_member(B)) &  | 
|
894  | 
(\<forall>A. subgroup_member(A) --> product(Gidentity::'a,A,A)) &  | 
|
895  | 
(\<forall>A. subgroup_member(A) --> product(A::'a,Gidentity,A)) &  | 
|
896  | 
(\<forall>A. subgroup_member(A) --> product(A::'a,Ginverse(A),Gidentity)) &  | 
|
897  | 
(\<forall>A. subgroup_member(A) --> product(Ginverse(A),A,Gidentity)) &  | 
|
898  | 
(\<forall>A. subgroup_member(A) --> subgroup_member(Ginverse(A))) &  | 
|
899  | 
(\<forall>A B. equal(A::'a,B) --> equal(Ginverse(A),Ginverse(B))) &  | 
|
900  | 
(\<forall>A C D B. product(A::'a,B,C) & product(A::'a,D,C) --> equal(D::'a,B)) &  | 
|
901  | 
(\<forall>B C D A. product(A::'a,B,C) & product(D::'a,B,C) --> equal(D::'a,A)) &  | 
|
902  | 
(subgroup_member(a)) &  | 
|
903  | 
(subgroup_member(Gidentity)) &  | 
|
| 24127 | 904  | 
(~equal(INVERSE(a),Ginverse(a))) --> False"  | 
905  | 
by meson  | 
|
906  | 
||
907  | 
(*163 inferences so far. Searching to depth 11. 0.3 secs*)  | 
|
908  | 
lemma GRP031_2:  | 
|
| 24128 | 909  | 
"(\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &  | 
910  | 
(\<forall>X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W)) &  | 
|
911  | 
(\<forall>Y U Z X V W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(U::'a,Z,W) --> product(X::'a,V,W)) &  | 
|
912  | 
(\<forall>Y X V U Z W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(X::'a,V,W) --> product(U::'a,Z,W)) &  | 
|
913  | 
(\<forall>A. product(A::'a,INVERSE(A),identity)) &  | 
|
914  | 
(\<forall>A. product(A::'a,identity,A)) &  | 
|
| 24127 | 915  | 
(\<forall>A. ~product(A::'a,a,identity)) --> False"  | 
916  | 
by meson  | 
|
917  | 
||
918  | 
(*47 inferences so far. Searching to depth 11. 0.2 secs*)  | 
|
919  | 
lemma GRP034_4:  | 
|
| 24128 | 920  | 
"(\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &  | 
921  | 
(\<forall>X. product(identity::'a,X,X)) &  | 
|
922  | 
(\<forall>X. product(X::'a,identity,X)) &  | 
|
923  | 
(\<forall>X. product(X::'a,INVERSE(X),identity)) &  | 
|
924  | 
(\<forall>Y U Z X V W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(U::'a,Z,W) --> product(X::'a,V,W)) &  | 
|
925  | 
(\<forall>Y X V U Z W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(X::'a,V,W) --> product(U::'a,Z,W)) &  | 
|
926  | 
(\<forall>B A C. subgroup_member(A) & subgroup_member(B) & product(B::'a,INVERSE(A),C) --> subgroup_member(C)) &  | 
|
927  | 
(subgroup_member(a)) &  | 
|
| 24127 | 928  | 
(~subgroup_member(INVERSE(a))) --> False"  | 
929  | 
by meson  | 
|
930  | 
||
931  | 
(*3287 inferences so far. Searching to depth 14. 3.5 secs*)  | 
|
932  | 
lemma GRP047_2:  | 
|
| 24128 | 933  | 
"(\<forall>X. product(identity::'a,X,X)) &  | 
934  | 
(\<forall>X. product(INVERSE(X),X,identity)) &  | 
|
935  | 
(\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &  | 
|
936  | 
(\<forall>X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W)) &  | 
|
937  | 
(\<forall>Y U Z X V W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(U::'a,Z,W) --> product(X::'a,V,W)) &  | 
|
938  | 
(\<forall>Y X V U Z W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(X::'a,V,W) --> product(U::'a,Z,W)) &  | 
|
939  | 
(\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) &  | 
|
940  | 
(equal(a::'a,b)) &  | 
|
| 24127 | 941  | 
(~equal(multiply(c::'a,a),multiply(c::'a,b))) --> False"  | 
942  | 
by meson  | 
|
943  | 
||
944  | 
(*25559 inferences so far. Searching to depth 19. 16.9 secs*)  | 
|
945  | 
lemma GRP130_1_002:  | 
|
| 24128 | 946  | 
"(group_element(e_1)) &  | 
947  | 
(group_element(e_2)) &  | 
|
948  | 
(~equal(e_1::'a,e_2)) &  | 
|
949  | 
(~equal(e_2::'a,e_1)) &  | 
|
950  | 
(\<forall>X Y. group_element(X) & group_element(Y) --> product(X::'a,Y,e_1) | product(X::'a,Y,e_2)) &  | 
|
951  | 
(\<forall>X Y W Z. product(X::'a,Y,W) & product(X::'a,Y,Z) --> equal(W::'a,Z)) &  | 
|
952  | 
(\<forall>X Y W Z. product(X::'a,W,Y) & product(X::'a,Z,Y) --> equal(W::'a,Z)) &  | 
|
953  | 
(\<forall>Y X W Z. product(W::'a,Y,X) & product(Z::'a,Y,X) --> equal(W::'a,Z)) &  | 
|
| 24127 | 954  | 
(\<forall>Z1 Z2 Y X. product(X::'a,Y,Z1) & product(X::'a,Z1,Z2) --> product(Z2::'a,Y,X)) --> False"  | 
955  | 
oops  | 
|
956  | 
||
957  | 
abbreviation "GRP004_0_ax INVERSE identity multiply equal \<equiv>  | 
|
| 24128 | 958  | 
(\<forall>X. equal(multiply(identity::'a,X),X)) &  | 
959  | 
(\<forall>X. equal(multiply(INVERSE(X),X),identity)) &  | 
|
960  | 
(\<forall>X Y Z. equal(multiply(multiply(X::'a,Y),Z),multiply(X::'a,multiply(Y::'a,Z)))) &  | 
|
961  | 
(\<forall>A B. equal(A::'a,B) --> equal(INVERSE(A),INVERSE(B))) &  | 
|
962  | 
(\<forall>C D E. equal(C::'a,D) --> equal(multiply(C::'a,E),multiply(D::'a,E))) &  | 
|
| 24127 | 963  | 
(\<forall>F' H G. equal(F'::'a,G) --> equal(multiply(H::'a,F'),multiply(H::'a,G)))"  | 
964  | 
||
965  | 
abbreviation "GRP004_2_ax multiply least_upper_bound greatest_lower_bound equal \<equiv>  | 
|
| 24128 | 966  | 
(\<forall>Y X. equal(greatest_lower_bound(X::'a,Y),greatest_lower_bound(Y::'a,X))) &  | 
967  | 
(\<forall>Y X. equal(least_upper_bound(X::'a,Y),least_upper_bound(Y::'a,X))) &  | 
|
968  | 
(\<forall>X Y Z. equal(greatest_lower_bound(X::'a,greatest_lower_bound(Y::'a,Z)),greatest_lower_bound(greatest_lower_bound(X::'a,Y),Z))) &  | 
|
969  | 
(\<forall>X Y Z. equal(least_upper_bound(X::'a,least_upper_bound(Y::'a,Z)),least_upper_bound(least_upper_bound(X::'a,Y),Z))) &  | 
|
970  | 
(\<forall>X. equal(least_upper_bound(X::'a,X),X)) &  | 
|
971  | 
(\<forall>X. equal(greatest_lower_bound(X::'a,X),X)) &  | 
|
972  | 
(\<forall>Y X. equal(least_upper_bound(X::'a,greatest_lower_bound(X::'a,Y)),X)) &  | 
|
973  | 
(\<forall>Y X. equal(greatest_lower_bound(X::'a,least_upper_bound(X::'a,Y)),X)) &  | 
|
974  | 
(\<forall>Y X Z. equal(multiply(X::'a,least_upper_bound(Y::'a,Z)),least_upper_bound(multiply(X::'a,Y),multiply(X::'a,Z)))) &  | 
|
975  | 
(\<forall>Y X Z. equal(multiply(X::'a,greatest_lower_bound(Y::'a,Z)),greatest_lower_bound(multiply(X::'a,Y),multiply(X::'a,Z)))) &  | 
|
976  | 
(\<forall>Y Z X. equal(multiply(least_upper_bound(Y::'a,Z),X),least_upper_bound(multiply(Y::'a,X),multiply(Z::'a,X)))) &  | 
|
977  | 
(\<forall>Y Z X. equal(multiply(greatest_lower_bound(Y::'a,Z),X),greatest_lower_bound(multiply(Y::'a,X),multiply(Z::'a,X)))) &  | 
|
978  | 
(\<forall>A B C. equal(A::'a,B) --> equal(greatest_lower_bound(A::'a,C),greatest_lower_bound(B::'a,C))) &  | 
|
979  | 
(\<forall>A C B. equal(A::'a,B) --> equal(greatest_lower_bound(C::'a,A),greatest_lower_bound(C::'a,B))) &  | 
|
980  | 
(\<forall>A B C. equal(A::'a,B) --> equal(least_upper_bound(A::'a,C),least_upper_bound(B::'a,C))) &  | 
|
981  | 
(\<forall>A C B. equal(A::'a,B) --> equal(least_upper_bound(C::'a,A),least_upper_bound(C::'a,B))) &  | 
|
982  | 
(\<forall>A B C. equal(A::'a,B) --> equal(multiply(A::'a,C),multiply(B::'a,C))) &  | 
|
| 24127 | 983  | 
(\<forall>A C B. equal(A::'a,B) --> equal(multiply(C::'a,A),multiply(C::'a,B)))"  | 
984  | 
||
985  | 
(*3468 inferences so far. Searching to depth 10. 9.1 secs*)  | 
|
986  | 
lemma GRP156_1:  | 
|
987  | 
"EQU001_0_ax equal &  | 
|
988  | 
GRP004_0_ax INVERSE identity multiply equal &  | 
|
989  | 
GRP004_2_ax multiply least_upper_bound greatest_lower_bound equal &  | 
|
| 24128 | 990  | 
(equal(least_upper_bound(a::'a,b),b)) &  | 
| 24127 | 991  | 
(~equal(greatest_lower_bound(multiply(a::'a,c),multiply(b::'a,c)),multiply(a::'a,c))) --> False"  | 
| 24128 | 992  | 
by meson  | 
| 24127 | 993  | 
|
994  | 
(*4394 inferences so far. Searching to depth 10. 8.2 secs*)  | 
|
995  | 
lemma GRP168_1:  | 
|
996  | 
"EQU001_0_ax equal &  | 
|
997  | 
GRP004_0_ax INVERSE identity multiply equal &  | 
|
| 24128 | 998  | 
GRP004_2_ax multiply least_upper_bound greatest_lower_bound equal &  | 
999  | 
(equal(least_upper_bound(a::'a,b),b)) &  | 
|
| 24127 | 1000  | 
(~equal(least_upper_bound(multiply(INVERSE(c),multiply(a::'a,c)),multiply(INVERSE(c),multiply(b::'a,c))),multiply(INVERSE(c),multiply(b::'a,c)))) --> False"  | 
1001  | 
by meson  | 
|
1002  | 
||
1003  | 
abbreviation "HEN002_0_ax identity Zero Divide equal mless_equal \<equiv>  | 
|
| 24128 | 1004  | 
(\<forall>X Y. mless_equal(X::'a,Y) --> equal(Divide(X::'a,Y),Zero)) &  | 
1005  | 
(\<forall>X Y. equal(Divide(X::'a,Y),Zero) --> mless_equal(X::'a,Y)) &  | 
|
1006  | 
(\<forall>Y X. mless_equal(Divide(X::'a,Y),X)) &  | 
|
1007  | 
(\<forall>X Y Z. mless_equal(Divide(Divide(X::'a,Z),Divide(Y::'a,Z)),Divide(Divide(X::'a,Y),Z))) &  | 
|
1008  | 
(\<forall>X. mless_equal(Zero::'a,X)) &  | 
|
1009  | 
(\<forall>X Y. mless_equal(X::'a,Y) & mless_equal(Y::'a,X) --> equal(X::'a,Y)) &  | 
|
| 24127 | 1010  | 
(\<forall>X. mless_equal(X::'a,identity))"  | 
1011  | 
||
1012  | 
abbreviation "HEN002_0_eq mless_equal Divide equal \<equiv>  | 
|
| 24128 | 1013  | 
(\<forall>A B C. equal(A::'a,B) --> equal(Divide(A::'a,C),Divide(B::'a,C))) &  | 
1014  | 
(\<forall>D F' E. equal(D::'a,E) --> equal(Divide(F'::'a,D),Divide(F'::'a,E))) &  | 
|
1015  | 
(\<forall>G H I'. equal(G::'a,H) & mless_equal(G::'a,I') --> mless_equal(H::'a,I')) &  | 
|
| 24127 | 1016  | 
(\<forall>J L K'. equal(J::'a,K') & mless_equal(L::'a,J) --> mless_equal(L::'a,K'))"  | 
1017  | 
||
1018  | 
(*250258 inferences so far. Searching to depth 16. 406.2 secs*)  | 
|
1019  | 
lemma HEN003_3:  | 
|
1020  | 
"EQU001_0_ax equal &  | 
|
1021  | 
HEN002_0_ax identity Zero Divide equal mless_equal &  | 
|
1022  | 
HEN002_0_eq mless_equal Divide equal &  | 
|
1023  | 
(~equal(Divide(a::'a,a),Zero)) --> False"  | 
|
1024  | 
oops  | 
|
1025  | 
||
1026  | 
(*202177 inferences so far. Searching to depth 14. 451 secs*)  | 
|
1027  | 
lemma HEN007_2:  | 
|
1028  | 
"EQU001_0_ax equal &  | 
|
| 24128 | 1029  | 
(\<forall>X Y. mless_equal(X::'a,Y) --> quotient(X::'a,Y,Zero)) &  | 
1030  | 
(\<forall>X Y. quotient(X::'a,Y,Zero) --> mless_equal(X::'a,Y)) &  | 
|
1031  | 
(\<forall>Y Z X. quotient(X::'a,Y,Z) --> mless_equal(Z::'a,X)) &  | 
|
1032  | 
(\<forall>Y X V3 V2 V1 Z V4 V5. quotient(X::'a,Y,V1) & quotient(Y::'a,Z,V2) & quotient(X::'a,Z,V3) & quotient(V3::'a,V2,V4) & quotient(V1::'a,Z,V5) --> mless_equal(V4::'a,V5)) &  | 
|
1033  | 
(\<forall>X. mless_equal(Zero::'a,X)) &  | 
|
1034  | 
(\<forall>X Y. mless_equal(X::'a,Y) & mless_equal(Y::'a,X) --> equal(X::'a,Y)) &  | 
|
1035  | 
(\<forall>X. mless_equal(X::'a,identity)) &  | 
|
1036  | 
(\<forall>X Y. quotient(X::'a,Y,Divide(X::'a,Y))) &  | 
|
1037  | 
(\<forall>X Y Z W. quotient(X::'a,Y,Z) & quotient(X::'a,Y,W) --> equal(Z::'a,W)) &  | 
|
1038  | 
(\<forall>X Y W Z. equal(X::'a,Y) & quotient(X::'a,W,Z) --> quotient(Y::'a,W,Z)) &  | 
|
1039  | 
(\<forall>X W Y Z. equal(X::'a,Y) & quotient(W::'a,X,Z) --> quotient(W::'a,Y,Z)) &  | 
|
1040  | 
(\<forall>X W Z Y. equal(X::'a,Y) & quotient(W::'a,Z,X) --> quotient(W::'a,Z,Y)) &  | 
|
1041  | 
(\<forall>X Z Y. equal(X::'a,Y) & mless_equal(Z::'a,X) --> mless_equal(Z::'a,Y)) &  | 
|
1042  | 
(\<forall>X Y Z. equal(X::'a,Y) & mless_equal(X::'a,Z) --> mless_equal(Y::'a,Z)) &  | 
|
1043  | 
(\<forall>X Y W. equal(X::'a,Y) --> equal(Divide(X::'a,W),Divide(Y::'a,W))) &  | 
|
1044  | 
(\<forall>X W Y. equal(X::'a,Y) --> equal(Divide(W::'a,X),Divide(W::'a,Y))) &  | 
|
1045  | 
(\<forall>X. quotient(X::'a,identity,Zero)) &  | 
|
1046  | 
(\<forall>X. quotient(Zero::'a,X,Zero)) &  | 
|
1047  | 
(\<forall>X. quotient(X::'a,X,Zero)) &  | 
|
1048  | 
(\<forall>X. quotient(X::'a,Zero,X)) &  | 
|
1049  | 
(\<forall>Y X Z. mless_equal(X::'a,Y) & mless_equal(Y::'a,Z) --> mless_equal(X::'a,Z)) &  | 
|
1050  | 
(\<forall>W1 X Z W2 Y. quotient(X::'a,Y,W1) & mless_equal(W1::'a,Z) & quotient(X::'a,Z,W2) --> mless_equal(W2::'a,Y)) &  | 
|
1051  | 
(mless_equal(x::'a,y)) &  | 
|
1052  | 
(quotient(z::'a,y,zQy)) &  | 
|
1053  | 
(quotient(z::'a,x,zQx)) &  | 
|
| 24127 | 1054  | 
(~mless_equal(zQy::'a,zQx)) --> False"  | 
1055  | 
oops  | 
|
1056  | 
||
1057  | 
(*60026 inferences so far. Searching to depth 12. 42.2 secs*)  | 
|
1058  | 
lemma HEN008_4:  | 
|
1059  | 
"EQU001_0_ax equal &  | 
|
1060  | 
HEN002_0_ax identity Zero Divide equal mless_equal &  | 
|
1061  | 
HEN002_0_eq mless_equal Divide equal &  | 
|
| 24128 | 1062  | 
(\<forall>X. equal(Divide(X::'a,identity),Zero)) &  | 
1063  | 
(\<forall>X. equal(Divide(Zero::'a,X),Zero)) &  | 
|
1064  | 
(\<forall>X. equal(Divide(X::'a,X),Zero)) &  | 
|
1065  | 
(equal(Divide(a::'a,Zero),a)) &  | 
|
1066  | 
(\<forall>Y X Z. mless_equal(X::'a,Y) & mless_equal(Y::'a,Z) --> mless_equal(X::'a,Z)) &  | 
|
1067  | 
(\<forall>X Z Y. mless_equal(Divide(X::'a,Y),Z) --> mless_equal(Divide(X::'a,Z),Y)) &  | 
|
1068  | 
(\<forall>Y Z X. mless_equal(X::'a,Y) --> mless_equal(Divide(Z::'a,Y),Divide(Z::'a,X))) &  | 
|
1069  | 
(mless_equal(a::'a,b)) &  | 
|
| 24127 | 1070  | 
(~mless_equal(Divide(a::'a,c),Divide(b::'a,c))) --> False"  | 
1071  | 
oops  | 
|
1072  | 
||
1073  | 
(*3160 inferences so far. Searching to depth 11. 3.5 secs*)  | 
|
1074  | 
lemma HEN009_5:  | 
|
1075  | 
"EQU001_0_ax equal &  | 
|
| 24128 | 1076  | 
(\<forall>Y X. equal(Divide(Divide(X::'a,Y),X),Zero)) &  | 
1077  | 
(\<forall>X Y Z. equal(Divide(Divide(Divide(X::'a,Z),Divide(Y::'a,Z)),Divide(Divide(X::'a,Y),Z)),Zero)) &  | 
|
1078  | 
(\<forall>X. equal(Divide(Zero::'a,X),Zero)) &  | 
|
1079  | 
(\<forall>X Y. equal(Divide(X::'a,Y),Zero) & equal(Divide(Y::'a,X),Zero) --> equal(X::'a,Y)) &  | 
|
1080  | 
(\<forall>X. equal(Divide(X::'a,identity),Zero)) &  | 
|
1081  | 
(\<forall>A B C. equal(A::'a,B) --> equal(Divide(A::'a,C),Divide(B::'a,C))) &  | 
|
1082  | 
(\<forall>D F' E. equal(D::'a,E) --> equal(Divide(F'::'a,D),Divide(F'::'a,E))) &  | 
|
1083  | 
(\<forall>Y X Z. equal(Divide(X::'a,Y),Zero) & equal(Divide(Y::'a,Z),Zero) --> equal(Divide(X::'a,Z),Zero)) &  | 
|
1084  | 
(\<forall>X Z Y. equal(Divide(Divide(X::'a,Y),Z),Zero) --> equal(Divide(Divide(X::'a,Z),Y),Zero)) &  | 
|
1085  | 
(\<forall>Y Z X. equal(Divide(X::'a,Y),Zero) --> equal(Divide(Divide(Z::'a,Y),Divide(Z::'a,X)),Zero)) &  | 
|
1086  | 
(~equal(Divide(identity::'a,a),Divide(identity::'a,Divide(identity::'a,Divide(identity::'a,a))))) &  | 
|
1087  | 
(equal(Divide(identity::'a,a),b)) &  | 
|
1088  | 
(equal(Divide(identity::'a,b),c)) &  | 
|
1089  | 
(equal(Divide(identity::'a,c),d)) &  | 
|
| 24127 | 1090  | 
(~equal(b::'a,d)) --> False"  | 
1091  | 
by meson  | 
|
1092  | 
||
1093  | 
(*970373 inferences so far. Searching to depth 17. 890.0 secs*)  | 
|
1094  | 
lemma HEN012_3:  | 
|
1095  | 
"EQU001_0_ax equal &  | 
|
1096  | 
HEN002_0_ax identity Zero Divide equal mless_equal &  | 
|
1097  | 
HEN002_0_eq mless_equal Divide equal &  | 
|
1098  | 
(~mless_equal(a::'a,a)) --> False"  | 
|
1099  | 
oops  | 
|
1100  | 
||
1101  | 
||
1102  | 
(*1063 inferences so far. Searching to depth 20. 1.0 secs*)  | 
|
1103  | 
lemma LCL010_1:  | 
|
| 24128 | 1104  | 
"(\<forall>X Y. is_a_theorem(equivalent(X::'a,Y)) & is_a_theorem(X) --> is_a_theorem(Y)) &  | 
1105  | 
(\<forall>X Z Y. is_a_theorem(equivalent(equivalent(X::'a,Y),equivalent(equivalent(X::'a,Z),equivalent(Z::'a,Y))))) &  | 
|
| 24127 | 1106  | 
(~is_a_theorem(equivalent(equivalent(a::'a,b),equivalent(equivalent(c::'a,b),equivalent(a::'a,c))))) --> False"  | 
1107  | 
by meson  | 
|
1108  | 
||
1109  | 
(*2549 inferences so far. Searching to depth 12. 1.4 secs*)  | 
|
1110  | 
lemma LCL077_2:  | 
|
| 24128 | 1111  | 
"(\<forall>X Y. is_a_theorem(implies(X,Y)) & is_a_theorem(X) --> is_a_theorem(Y)) &  | 
1112  | 
(\<forall>Y X. is_a_theorem(implies(X,implies(Y,X)))) &  | 
|
1113  | 
(\<forall>Y X Z. is_a_theorem(implies(implies(X,implies(Y,Z)),implies(implies(X,Y),implies(X,Z))))) &  | 
|
1114  | 
(\<forall>Y X. is_a_theorem(implies(implies(not(X),not(Y)),implies(Y,X)))) &  | 
|
1115  | 
(\<forall>X2 X1 X3. is_a_theorem(implies(X1,X2)) & is_a_theorem(implies(X2,X3)) --> is_a_theorem(implies(X1,X3))) &  | 
|
| 24127 | 1116  | 
(~is_a_theorem(implies(not(not(a)),a))) --> False"  | 
1117  | 
by meson  | 
|
1118  | 
||
1119  | 
(*2036 inferences so far. Searching to depth 20. 1.5 secs*)  | 
|
1120  | 
lemma LCL082_1:  | 
|
| 24128 | 1121  | 
"(\<forall>X Y. is_a_theorem(implies(X::'a,Y)) & is_a_theorem(X) --> is_a_theorem(Y)) &  | 
1122  | 
(\<forall>Y Z U X. is_a_theorem(implies(implies(implies(X::'a,Y),Z),implies(implies(Z::'a,X),implies(U::'a,X))))) &  | 
|
| 24127 | 1123  | 
(~is_a_theorem(implies(a::'a,implies(b::'a,a)))) --> False"  | 
1124  | 
by meson  | 
|
1125  | 
||
1126  | 
(*1100 inferences so far. Searching to depth 13. 1.0 secs*)  | 
|
1127  | 
lemma LCL111_1:  | 
|
| 24128 | 1128  | 
"(\<forall>X Y. is_a_theorem(implies(X,Y)) & is_a_theorem(X) --> is_a_theorem(Y)) &  | 
1129  | 
(\<forall>Y X. is_a_theorem(implies(X,implies(Y,X)))) &  | 
|
1130  | 
(\<forall>Y X Z. is_a_theorem(implies(implies(X,Y),implies(implies(Y,Z),implies(X,Z))))) &  | 
|
1131  | 
(\<forall>Y X. is_a_theorem(implies(implies(implies(X,Y),Y),implies(implies(Y,X),X)))) &  | 
|
1132  | 
(\<forall>Y X. is_a_theorem(implies(implies(not(X),not(Y)),implies(Y,X)))) &  | 
|
| 24127 | 1133  | 
(~is_a_theorem(implies(implies(a,b),implies(implies(c,a),implies(c,b))))) --> False"  | 
1134  | 
by meson  | 
|
1135  | 
||
1136  | 
(*667 inferences so far. Searching to depth 9. 1.4 secs*)  | 
|
1137  | 
lemma LCL143_1:  | 
|
| 24128 | 1138  | 
"(\<forall>X. equal(X,X)) &  | 
1139  | 
(\<forall>Y X. equal(X,Y) --> equal(Y,X)) &  | 
|
1140  | 
(\<forall>Y X Z. equal(X,Y) & equal(Y,Z) --> equal(X,Z)) &  | 
|
1141  | 
(\<forall>X. equal(implies(true,X),X)) &  | 
|
1142  | 
(\<forall>Y X Z. equal(implies(implies(X,Y),implies(implies(Y,Z),implies(X,Z))),true)) &  | 
|
1143  | 
(\<forall>Y X. equal(implies(implies(X,Y),Y),implies(implies(Y,X),X))) &  | 
|
1144  | 
(\<forall>Y X. equal(implies(implies(not(X),not(Y)),implies(Y,X)),true)) &  | 
|
1145  | 
(\<forall>A B C. equal(A,B) --> equal(implies(A,C),implies(B,C))) &  | 
|
1146  | 
(\<forall>D F' E. equal(D,E) --> equal(implies(F',D),implies(F',E))) &  | 
|
1147  | 
(\<forall>G H. equal(G,H) --> equal(not(G),not(H))) &  | 
|
1148  | 
(\<forall>X Y. equal(big_V(X,Y),implies(implies(X,Y),Y))) &  | 
|
1149  | 
(\<forall>X Y. equal(big_hat(X,Y),not(big_V(not(X),not(Y))))) &  | 
|
1150  | 
(\<forall>X Y. ordered(X,Y) --> equal(implies(X,Y),true)) &  | 
|
1151  | 
(\<forall>X Y. equal(implies(X,Y),true) --> ordered(X,Y)) &  | 
|
1152  | 
(\<forall>A B C. equal(A,B) --> equal(big_V(A,C),big_V(B,C))) &  | 
|
1153  | 
(\<forall>D F' E. equal(D,E) --> equal(big_V(F',D),big_V(F',E))) &  | 
|
1154  | 
(\<forall>G H I'. equal(G,H) --> equal(big_hat(G,I'),big_hat(H,I'))) &  | 
|
1155  | 
(\<forall>J L K'. equal(J,K') --> equal(big_hat(L,J),big_hat(L,K'))) &  | 
|
1156  | 
(\<forall>M N O'. equal(M,N) & ordered(M,O') --> ordered(N,O')) &  | 
|
1157  | 
(\<forall>P R Q. equal(P,Q) & ordered(R,P) --> ordered(R,Q)) &  | 
|
1158  | 
(ordered(x,y)) &  | 
|
| 24127 | 1159  | 
(~ordered(implies(z,x),implies(z,y))) --> False"  | 
1160  | 
by meson  | 
|
1161  | 
||
1162  | 
(*5245 inferences so far. Searching to depth 12. 4.6 secs*)  | 
|
1163  | 
lemma LCL182_1:  | 
|
| 24128 | 1164  | 
"(\<forall>A. axiom(or(not(or(A,A)),A))) &  | 
1165  | 
(\<forall>B A. axiom(or(not(A),or(B,A)))) &  | 
|
1166  | 
(\<forall>B A. axiom(or(not(or(A,B)),or(B,A)))) &  | 
|
1167  | 
(\<forall>B A C. axiom(or(not(or(A,or(B,C))),or(B,or(A,C))))) &  | 
|
1168  | 
(\<forall>A C B. axiom(or(not(or(not(A),B)),or(not(or(C,A)),or(C,B))))) &  | 
|
1169  | 
(\<forall>X. axiom(X) --> theorem(X)) &  | 
|
1170  | 
(\<forall>X Y. axiom(or(not(Y),X)) & theorem(Y) --> theorem(X)) &  | 
|
1171  | 
(\<forall>X Y Z. axiom(or(not(X),Y)) & theorem(or(not(Y),Z)) --> theorem(or(not(X),Z))) &  | 
|
| 24127 | 1172  | 
(~theorem(or(not(or(not(p),q)),or(not(not(q)),not(p))))) --> False"  | 
1173  | 
by meson  | 
|
1174  | 
||
1175  | 
(*410 inferences so far. Searching to depth 10. 0.3 secs*)  | 
|
1176  | 
lemma LCL200_1:  | 
|
| 24128 | 1177  | 
"(\<forall>A. axiom(or(not(or(A,A)),A))) &  | 
1178  | 
(\<forall>B A. axiom(or(not(A),or(B,A)))) &  | 
|
1179  | 
(\<forall>B A. axiom(or(not(or(A,B)),or(B,A)))) &  | 
|
1180  | 
(\<forall>B A C. axiom(or(not(or(A,or(B,C))),or(B,or(A,C))))) &  | 
|
1181  | 
(\<forall>A C B. axiom(or(not(or(not(A),B)),or(not(or(C,A)),or(C,B))))) &  | 
|
1182  | 
(\<forall>X. axiom(X) --> theorem(X)) &  | 
|
1183  | 
(\<forall>X Y. axiom(or(not(Y),X)) & theorem(Y) --> theorem(X)) &  | 
|
1184  | 
(\<forall>X Y Z. axiom(or(not(X),Y)) & theorem(or(not(Y),Z)) --> theorem(or(not(X),Z))) &  | 
|
| 24127 | 1185  | 
(~theorem(or(not(not(or(p,q))),not(q)))) --> False"  | 
1186  | 
by meson  | 
|
1187  | 
||
1188  | 
(*5849 inferences so far. Searching to depth 12. 5.6 secs*)  | 
|
1189  | 
lemma LCL215_1:  | 
|
| 24128 | 1190  | 
"(\<forall>A. axiom(or(not(or(A,A)),A))) &  | 
1191  | 
(\<forall>B A. axiom(or(not(A),or(B,A)))) &  | 
|
1192  | 
(\<forall>B A. axiom(or(not(or(A,B)),or(B,A)))) &  | 
|
1193  | 
(\<forall>B A C. axiom(or(not(or(A,or(B,C))),or(B,or(A,C))))) &  | 
|
1194  | 
(\<forall>A C B. axiom(or(not(or(not(A),B)),or(not(or(C,A)),or(C,B))))) &  | 
|
1195  | 
(\<forall>X. axiom(X) --> theorem(X)) &  | 
|
1196  | 
(\<forall>X Y. axiom(or(not(Y),X)) & theorem(Y) --> theorem(X)) &  | 
|
1197  | 
(\<forall>X Y Z. axiom(or(not(X),Y)) & theorem(or(not(Y),Z)) --> theorem(or(not(X),Z))) &  | 
|
| 24127 | 1198  | 
(~theorem(or(not(or(not(p),q)),or(not(or(p,q)),q)))) --> False"  | 
1199  | 
by meson  | 
|
1200  | 
||
1201  | 
(*0 secs. Not sure that a search even starts!*)  | 
|
1202  | 
lemma LCL230_2:  | 
|
| 24128 | 1203  | 
"(q --> p | r) &  | 
1204  | 
(~p) &  | 
|
1205  | 
(q) &  | 
|
| 24127 | 1206  | 
(~r) --> False"  | 
1207  | 
by meson  | 
|
1208  | 
||
1209  | 
(*119585 inferences so far. Searching to depth 14. 262.4 secs*)  | 
|
1210  | 
lemma LDA003_1:  | 
|
1211  | 
"EQU001_0_ax equal &  | 
|
| 24128 | 1212  | 
(\<forall>Y X Z. equal(f(X::'a,f(Y::'a,Z)),f(f(X::'a,Y),f(X::'a,Z)))) &  | 
1213  | 
(\<forall>X Y. left(X::'a,f(X::'a,Y))) &  | 
|
1214  | 
(\<forall>Y X Z. left(X::'a,Y) & left(Y::'a,Z) --> left(X::'a,Z)) &  | 
|
1215  | 
(equal(num2::'a,f(num1::'a,num1))) &  | 
|
1216  | 
(equal(num3::'a,f(num2::'a,num1))) &  | 
|
1217  | 
(equal(u::'a,f(num2::'a,num2))) &  | 
|
1218  | 
(\<forall>A B C. equal(A::'a,B) --> equal(f(A::'a,C),f(B::'a,C))) &  | 
|
1219  | 
(\<forall>D F' E. equal(D::'a,E) --> equal(f(F'::'a,D),f(F'::'a,E))) &  | 
|
1220  | 
(\<forall>G H I'. equal(G::'a,H) & left(G::'a,I') --> left(H::'a,I')) &  | 
|
1221  | 
(\<forall>J L K'. equal(J::'a,K') & left(L::'a,J) --> left(L::'a,K')) &  | 
|
| 24127 | 1222  | 
(~left(num3::'a,u)) --> False"  | 
1223  | 
oops  | 
|
1224  | 
||
1225  | 
||
1226  | 
(*2392 inferences so far. Searching to depth 12. 2.2 secs*)  | 
|
1227  | 
lemma MSC002_1:  | 
|
| 24128 | 1228  | 
"(at(something::'a,here,now)) &  | 
1229  | 
(\<forall>Place Situation. hand_at(Place::'a,Situation) --> hand_at(Place::'a,let_go(Situation))) &  | 
|
1230  | 
(\<forall>Place Another_place Situation. hand_at(Place::'a,Situation) --> hand_at(Another_place::'a,go(Another_place::'a,Situation))) &  | 
|
1231  | 
(\<forall>Thing Situation. ~held(Thing::'a,let_go(Situation))) &  | 
|
1232  | 
(\<forall>Situation Thing. at(Thing::'a,here,Situation) --> red(Thing)) &  | 
|
1233  | 
(\<forall>Thing Place Situation. at(Thing::'a,Place,Situation) --> at(Thing::'a,Place,let_go(Situation))) &  | 
|
1234  | 
(\<forall>Thing Place Situation. at(Thing::'a,Place,Situation) --> at(Thing::'a,Place,pick_up(Situation))) &  | 
|
1235  | 
(\<forall>Thing Place Situation. at(Thing::'a,Place,Situation) --> grabbed(Thing::'a,pick_up(go(Place::'a,let_go(Situation))))) &  | 
|
1236  | 
(\<forall>Thing Situation. red(Thing) & put(Thing::'a,there,Situation) --> answer(Situation)) &  | 
|
1237  | 
(\<forall>Place Thing Another_place Situation. at(Thing::'a,Place,Situation) & grabbed(Thing::'a,Situation) --> put(Thing::'a,Another_place,go(Another_place::'a,Situation))) &  | 
|
1238  | 
(\<forall>Thing Place Another_place Situation. at(Thing::'a,Place,Situation) --> held(Thing::'a,Situation) | at(Thing::'a,Place,go(Another_place::'a,Situation))) &  | 
|
1239  | 
(\<forall>One_place Thing Place Situation. hand_at(One_place::'a,Situation) & held(Thing::'a,Situation) --> at(Thing::'a,Place,go(Place::'a,Situation))) &  | 
|
1240  | 
(\<forall>Place Thing Situation. hand_at(Place::'a,Situation) & at(Thing::'a,Place,Situation) --> held(Thing::'a,pick_up(Situation))) &  | 
|
| 24127 | 1241  | 
(\<forall>Situation. ~answer(Situation)) --> False"  | 
1242  | 
by meson  | 
|
1243  | 
||
1244  | 
(*73 inferences so far. Searching to depth 10. 0.2 secs*)  | 
|
1245  | 
lemma MSC003_1:  | 
|
| 24128 | 1246  | 
"(\<forall>Number_of_small_parts Small_part Big_part Number_of_mid_parts Mid_part. has_parts(Big_part::'a,Number_of_mid_parts,Mid_part) --> in'(object_in(Big_part::'a,Mid_part,Small_part,Number_of_mid_parts,Number_of_small_parts),Mid_part) | has_parts(Big_part::'a,mtimes(Number_of_mid_parts::'a,Number_of_small_parts),Small_part)) &  | 
1247  | 
(\<forall>Big_part Mid_part Number_of_mid_parts Number_of_small_parts Small_part. has_parts(Big_part::'a,Number_of_mid_parts,Mid_part) & has_parts(object_in(Big_part::'a,Mid_part,Small_part,Number_of_mid_parts,Number_of_small_parts),Number_of_small_parts,Small_part) --> has_parts(Big_part::'a,mtimes(Number_of_mid_parts::'a,Number_of_small_parts),Small_part)) &  | 
|
1248  | 
(in'(john::'a,boy)) &  | 
|
1249  | 
(\<forall>X. in'(X::'a,boy) --> in'(X::'a,human)) &  | 
|
1250  | 
(\<forall>X. in'(X::'a,hand) --> has_parts(X::'a,num5,fingers)) &  | 
|
1251  | 
(\<forall>X. in'(X::'a,human) --> has_parts(X::'a,num2,arm)) &  | 
|
1252  | 
(\<forall>X. in'(X::'a,arm) --> has_parts(X::'a,num1,hand)) &  | 
|
| 24127 | 1253  | 
(~has_parts(john::'a,mtimes(num2::'a,num1),hand)) --> False"  | 
1254  | 
by meson  | 
|
1255  | 
||
1256  | 
(*1486 inferences so far. Searching to depth 20. 1.2 secs*)  | 
|
1257  | 
lemma MSC004_1:  | 
|
| 24128 | 1258  | 
"(\<forall>Number_of_small_parts Small_part Big_part Number_of_mid_parts Mid_part. has_parts(Big_part::'a,Number_of_mid_parts,Mid_part) --> in'(object_in(Big_part::'a,Mid_part,Small_part,Number_of_mid_parts,Number_of_small_parts),Mid_part) | has_parts(Big_part::'a,mtimes(Number_of_mid_parts::'a,Number_of_small_parts),Small_part)) &  | 
1259  | 
(\<forall>Big_part Mid_part Number_of_mid_parts Number_of_small_parts Small_part. has_parts(Big_part::'a,Number_of_mid_parts,Mid_part) & has_parts(object_in(Big_part::'a,Mid_part,Small_part,Number_of_mid_parts,Number_of_small_parts),Number_of_small_parts,Small_part) --> has_parts(Big_part::'a,mtimes(Number_of_mid_parts::'a,Number_of_small_parts),Small_part)) &  | 
|
1260  | 
(in'(john::'a,boy)) &  | 
|
1261  | 
(\<forall>X. in'(X::'a,boy) --> in'(X::'a,human)) &  | 
|
1262  | 
(\<forall>X. in'(X::'a,hand) --> has_parts(X::'a,num5,fingers)) &  | 
|
1263  | 
(\<forall>X. in'(X::'a,human) --> has_parts(X::'a,num2,arm)) &  | 
|
1264  | 
(\<forall>X. in'(X::'a,arm) --> has_parts(X::'a,num1,hand)) &  | 
|
| 24127 | 1265  | 
(~has_parts(john::'a,mtimes(mtimes(num2::'a,num1),num5),fingers)) --> False"  | 
1266  | 
by meson  | 
|
1267  | 
||
1268  | 
(*100 inferences so far. Searching to depth 12. 0.1 secs*)  | 
|
1269  | 
lemma MSC005_1:  | 
|
| 24128 | 1270  | 
"(value(truth::'a,truth)) &  | 
1271  | 
(value(falsity::'a,falsity)) &  | 
|
1272  | 
(\<forall>X Y. value(X::'a,truth) & value(Y::'a,truth) --> value(xor(X::'a,Y),falsity)) &  | 
|
1273  | 
(\<forall>X Y. value(X::'a,truth) & value(Y::'a,falsity) --> value(xor(X::'a,Y),truth)) &  | 
|
1274  | 
(\<forall>X Y. value(X::'a,falsity) & value(Y::'a,truth) --> value(xor(X::'a,Y),truth)) &  | 
|
1275  | 
(\<forall>X Y. value(X::'a,falsity) & value(Y::'a,falsity) --> value(xor(X::'a,Y),falsity)) &  | 
|
| 24127 | 1276  | 
(\<forall>Value. ~value(xor(xor(xor(xor(truth::'a,falsity),falsity),truth),falsity),Value)) --> False"  | 
1277  | 
by meson  | 
|
1278  | 
||
1279  | 
(*19116 inferences so far. Searching to depth 16. 15.9 secs*)  | 
|
1280  | 
lemma MSC006_1:  | 
|
| 24128 | 1281  | 
"(\<forall>Y X Z. p(X::'a,Y) & p(Y::'a,Z) --> p(X::'a,Z)) &  | 
1282  | 
(\<forall>Y X Z. q(X::'a,Y) & q(Y::'a,Z) --> q(X::'a,Z)) &  | 
|
1283  | 
(\<forall>Y X. q(X::'a,Y) --> q(Y::'a,X)) &  | 
|
1284  | 
(\<forall>X Y. p(X::'a,Y) | q(X::'a,Y)) &  | 
|
1285  | 
(~p(a::'a,b)) &  | 
|
| 24127 | 1286  | 
(~q(c::'a,d)) --> False"  | 
1287  | 
by meson  | 
|
1288  | 
||
1289  | 
(*1713 inferences so far. Searching to depth 10. 2.8 secs*)  | 
|
1290  | 
lemma NUM001_1:  | 
|
| 24128 | 1291  | 
"(\<forall>A. equal(A::'a,A)) &  | 
1292  | 
(\<forall>B A C. equal(A::'a,B) & equal(B::'a,C) --> equal(A::'a,C)) &  | 
|
1293  | 
(\<forall>B A. equal(add(A::'a,B),add(B::'a,A))) &  | 
|
1294  | 
(\<forall>A B C. equal(add(A::'a,add(B::'a,C)),add(add(A::'a,B),C))) &  | 
|
1295  | 
(\<forall>B A. equal(subtract(add(A::'a,B),B),A)) &  | 
|
1296  | 
(\<forall>A B. equal(A::'a,subtract(add(A::'a,B),B))) &  | 
|
1297  | 
(\<forall>A C B. equal(add(subtract(A::'a,B),C),subtract(add(A::'a,C),B))) &  | 
|
1298  | 
(\<forall>A C B. equal(subtract(add(A::'a,B),C),add(subtract(A::'a,C),B))) &  | 
|
1299  | 
(\<forall>A C B D. equal(A::'a,B) & equal(C::'a,add(A::'a,D)) --> equal(C::'a,add(B::'a,D))) &  | 
|
1300  | 
(\<forall>A C D B. equal(A::'a,B) & equal(C::'a,add(D::'a,A)) --> equal(C::'a,add(D::'a,B))) &  | 
|
1301  | 
(\<forall>A C B D. equal(A::'a,B) & equal(C::'a,subtract(A::'a,D)) --> equal(C::'a,subtract(B::'a,D))) &  | 
|
1302  | 
(\<forall>A C D B. equal(A::'a,B) & equal(C::'a,subtract(D::'a,A)) --> equal(C::'a,subtract(D::'a,B))) &  | 
|
| 24127 | 1303  | 
(~equal(add(add(a::'a,b),c),add(a::'a,add(b::'a,c)))) --> False"  | 
1304  | 
by meson  | 
|
1305  | 
||
1306  | 
abbreviation "NUM001_0_ax multiply successor num0 add equal \<equiv>  | 
|
| 24128 | 1307  | 
(\<forall>A. equal(add(A::'a,num0),A)) &  | 
1308  | 
(\<forall>A B. equal(add(A::'a,successor(B)),successor(add(A::'a,B)))) &  | 
|
1309  | 
(\<forall>A. equal(multiply(A::'a,num0),num0)) &  | 
|
1310  | 
(\<forall>B A. equal(multiply(A::'a,successor(B)),add(multiply(A::'a,B),A))) &  | 
|
1311  | 
(\<forall>A B. equal(successor(A),successor(B)) --> equal(A::'a,B)) &  | 
|
| 24127 | 1312  | 
(\<forall>A B. equal(A::'a,B) --> equal(successor(A),successor(B)))"  | 
1313  | 
||
1314  | 
abbreviation "NUM001_1_ax predecessor_of_1st_minus_2nd successor add equal mless \<equiv>  | 
|
| 24128 | 1315  | 
(\<forall>A C B. mless(A::'a,B) & mless(C::'a,A) --> mless(C::'a,B)) &  | 
1316  | 
(\<forall>A B C. equal(add(successor(A),B),C) --> mless(B::'a,C)) &  | 
|
| 24127 | 1317  | 
(\<forall>A B. mless(A::'a,B) --> equal(add(successor(predecessor_of_1st_minus_2nd(B::'a,A)),A),B))"  | 
1318  | 
||
1319  | 
abbreviation "NUM001_2_ax equal mless divides \<equiv>  | 
|
| 24128 | 1320  | 
(\<forall>A B. divides(A::'a,B) --> mless(A::'a,B) | equal(A::'a,B)) &  | 
1321  | 
(\<forall>A B. mless(A::'a,B) --> divides(A::'a,B)) &  | 
|
| 24127 | 1322  | 
(\<forall>A B. equal(A::'a,B) --> divides(A::'a,B))"  | 
1323  | 
||
1324  | 
(*20717 inferences so far. Searching to depth 11. 13.7 secs*)  | 
|
1325  | 
lemma NUM021_1:  | 
|
1326  | 
"EQU001_0_ax equal &  | 
|
1327  | 
NUM001_0_ax multiply successor num0 add equal &  | 
|
1328  | 
NUM001_1_ax predecessor_of_1st_minus_2nd successor add equal mless &  | 
|
1329  | 
NUM001_2_ax equal mless divides &  | 
|
| 24128 | 1330  | 
(mless(b::'a,c)) &  | 
1331  | 
(~mless(b::'a,a)) &  | 
|
1332  | 
(divides(c::'a,a)) &  | 
|
| 24127 | 1333  | 
(\<forall>A. ~equal(successor(A),num0)) --> False"  | 
1334  | 
by meson  | 
|
1335  | 
||
1336  | 
(*26320 inferences so far. Searching to depth 10. 26.4 secs*)  | 
|
1337  | 
lemma NUM024_1:  | 
|
1338  | 
"EQU001_0_ax equal &  | 
|
1339  | 
NUM001_0_ax multiply successor num0 add equal &  | 
|
| 24128 | 1340  | 
NUM001_1_ax predecessor_of_1st_minus_2nd successor add equal mless &  | 
1341  | 
(\<forall>B A. equal(add(A::'a,B),add(B::'a,A))) &  | 
|
1342  | 
(\<forall>B A C. equal(add(A::'a,B),add(C::'a,B)) --> equal(A::'a,C)) &  | 
|
1343  | 
(mless(a::'a,a)) &  | 
|
| 24127 | 1344  | 
(\<forall>A. ~equal(successor(A),num0)) --> False"  | 
1345  | 
oops  | 
|
1346  | 
||
1347  | 
abbreviation "SET004_0_ax not_homomorphism2 not_homomorphism1  | 
|
1348  | 
homomorphism compatible operation cantor diagonalise subset_relation  | 
|
1349  | 
one_to_one choice apply regular function identity_relation  | 
|
1350  | 
single_valued_class compos powerClass sum_class omega inductive  | 
|
1351  | 
successor_relation successor image' rng domain range_of INVERSE flip  | 
|
1352  | 
rot domain_of null_class restrct difference union complement  | 
|
1353  | 
intersection element_relation second first cross_product ordered_pair  | 
|
1354  | 
singleton unordered_pair equal universal_class not_subclass_element  | 
|
1355  | 
member subclass \<equiv>  | 
|
| 24128 | 1356  | 
(\<forall>X U Y. subclass(X::'a,Y) & member(U::'a,X) --> member(U::'a,Y)) &  | 
1357  | 
(\<forall>X Y. member(not_subclass_element(X::'a,Y),X) | subclass(X::'a,Y)) &  | 
|
1358  | 
(\<forall>X Y. member(not_subclass_element(X::'a,Y),Y) --> subclass(X::'a,Y)) &  | 
|
1359  | 
(\<forall>X. subclass(X::'a,universal_class)) &  | 
|
1360  | 
(\<forall>X Y. equal(X::'a,Y) --> subclass(X::'a,Y)) &  | 
|
1361  | 
(\<forall>Y X. equal(X::'a,Y) --> subclass(Y::'a,X)) &  | 
|
1362  | 
(\<forall>X Y. subclass(X::'a,Y) & subclass(Y::'a,X) --> equal(X::'a,Y)) &  | 
|
1363  | 
(\<forall>X U Y. member(U::'a,unordered_pair(X::'a,Y)) --> equal(U::'a,X) | equal(U::'a,Y)) &  | 
|
1364  | 
(\<forall>X Y. member(X::'a,universal_class) --> member(X::'a,unordered_pair(X::'a,Y))) &  | 
|
1365  | 
(\<forall>X Y. member(Y::'a,universal_class) --> member(Y::'a,unordered_pair(X::'a,Y))) &  | 
|
1366  | 
(\<forall>X Y. member(unordered_pair(X::'a,Y),universal_class)) &  | 
|
1367  | 
(\<forall>X. equal(unordered_pair(X::'a,X),singleton(X))) &  | 
|
1368  | 
(\<forall>X Y. equal(unordered_pair(singleton(X),unordered_pair(X::'a,singleton(Y))),ordered_pair(X::'a,Y))) &  | 
|
1369  | 
(\<forall>V Y U X. member(ordered_pair(U::'a,V),cross_product(X::'a,Y)) --> member(U::'a,X)) &  | 
|
1370  | 
(\<forall>U X V Y. member(ordered_pair(U::'a,V),cross_product(X::'a,Y)) --> member(V::'a,Y)) &  | 
|
1371  | 
(\<forall>U V X Y. member(U::'a,X) & member(V::'a,Y) --> member(ordered_pair(U::'a,V),cross_product(X::'a,Y))) &  | 
|
1372  | 
(\<forall>X Y Z. member(Z::'a,cross_product(X::'a,Y)) --> equal(ordered_pair(first(Z),second(Z)),Z)) &  | 
|
1373  | 
(subclass(element_relation::'a,cross_product(universal_class::'a,universal_class))) &  | 
|
1374  | 
(\<forall>X Y. member(ordered_pair(X::'a,Y),element_relation) --> member(X::'a,Y)) &  | 
|
1375  | 
(\<forall>X Y. member(ordered_pair(X::'a,Y),cross_product(universal_class::'a,universal_class)) & member(X::'a,Y) --> member(ordered_pair(X::'a,Y),element_relation)) &  | 
|
1376  | 
(\<forall>Y Z X. member(Z::'a,intersection(X::'a,Y)) --> member(Z::'a,X)) &  | 
|
1377  | 
(\<forall>X Z Y. member(Z::'a,intersection(X::'a,Y)) --> member(Z::'a,Y)) &  | 
|
1378  | 
(\<forall>Z X Y. member(Z::'a,X) & member(Z::'a,Y) --> member(Z::'a,intersection(X::'a,Y))) &  | 
|
1379  | 
(\<forall>Z X. ~(member(Z::'a,complement(X)) & member(Z::'a,X))) &  | 
|
1380  | 
(\<forall>Z X. member(Z::'a,universal_class) --> member(Z::'a,complement(X)) | member(Z::'a,X)) &  | 
|
1381  | 
(\<forall>X Y. equal(complement(intersection(complement(X),complement(Y))),union(X::'a,Y))) &  | 
|
1382  | 
(\<forall>X Y. equal(intersection(complement(intersection(X::'a,Y)),complement(intersection(complement(X),complement(Y)))),difference(X::'a,Y))) &  | 
|
1383  | 
(\<forall>Xr X Y. equal(intersection(Xr::'a,cross_product(X::'a,Y)),restrct(Xr::'a,X,Y))) &  | 
|
1384  | 
(\<forall>Xr X Y. equal(intersection(cross_product(X::'a,Y),Xr),restrct(Xr::'a,X,Y))) &  | 
|
1385  | 
(\<forall>Z X. ~(equal(restrct(X::'a,singleton(Z),universal_class),null_class) & member(Z::'a,domain_of(X)))) &  | 
|
1386  | 
(\<forall>Z X. member(Z::'a,universal_class) --> equal(restrct(X::'a,singleton(Z),universal_class),null_class) | member(Z::'a,domain_of(X))) &  | 
|
1387  | 
(\<forall>X. subclass(rot(X),cross_product(cross_product(universal_class::'a,universal_class),universal_class))) &  | 
|
1388  | 
(\<forall>V W U X. member(ordered_pair(ordered_pair(U::'a,V),W),rot(X)) --> member(ordered_pair(ordered_pair(V::'a,W),U),X)) &  | 
|
1389  | 
(\<forall>U V W X. member(ordered_pair(ordered_pair(V::'a,W),U),X) & member(ordered_pair(ordered_pair(U::'a,V),W),cross_product(cross_product(universal_class::'a,universal_class),universal_class)) --> member(ordered_pair(ordered_pair(U::'a,V),W),rot(X))) &  | 
|
1390  | 
(\<forall>X. subclass(flip(X),cross_product(cross_product(universal_class::'a,universal_class),universal_class))) &  | 
|
1391  | 
(\<forall>V U W X. member(ordered_pair(ordered_pair(U::'a,V),W),flip(X)) --> member(ordered_pair(ordered_pair(V::'a,U),W),X)) &  | 
|
1392  | 
(\<forall>U V W X. member(ordered_pair(ordered_pair(V::'a,U),W),X) & member(ordered_pair(ordered_pair(U::'a,V),W),cross_product(cross_product(universal_class::'a,universal_class),universal_class)) --> member(ordered_pair(ordered_pair(U::'a,V),W),flip(X))) &  | 
|
1393  | 
(\<forall>Y. equal(domain_of(flip(cross_product(Y::'a,universal_class))),INVERSE(Y))) &  | 
|
1394  | 
(\<forall>Z. equal(domain_of(INVERSE(Z)),range_of(Z))) &  | 
|
1395  | 
(\<forall>Z X Y. equal(first(not_subclass_element(restrct(Z::'a,X,singleton(Y)),null_class)),domain(Z::'a,X,Y))) &  | 
|
1396  | 
(\<forall>Z X Y. equal(second(not_subclass_element(restrct(Z::'a,singleton(X),Y),null_class)),rng(Z::'a,X,Y))) &  | 
|
1397  | 
(\<forall>Xr X. equal(range_of(restrct(Xr::'a,X,universal_class)),image'(Xr::'a,X))) &  | 
|
1398  | 
(\<forall>X. equal(union(X::'a,singleton(X)),successor(X))) &  | 
|
1399  | 
(subclass(successor_relation::'a,cross_product(universal_class::'a,universal_class))) &  | 
|
1400  | 
(\<forall>X Y. member(ordered_pair(X::'a,Y),successor_relation) --> equal(successor(X),Y)) &  | 
|
1401  | 
(\<forall>X Y. equal(successor(X),Y) & member(ordered_pair(X::'a,Y),cross_product(universal_class::'a,universal_class)) --> member(ordered_pair(X::'a,Y),successor_relation)) &  | 
|
1402  | 
(\<forall>X. inductive(X) --> member(null_class::'a,X)) &  | 
|
1403  | 
(\<forall>X. inductive(X) --> subclass(image'(successor_relation::'a,X),X)) &  | 
|
1404  | 
(\<forall>X. member(null_class::'a,X) & subclass(image'(successor_relation::'a,X),X) --> inductive(X)) &  | 
|
1405  | 
(inductive(omega)) &  | 
|
1406  | 
(\<forall>Y. inductive(Y) --> subclass(omega::'a,Y)) &  | 
|
1407  | 
(member(omega::'a,universal_class)) &  | 
|
1408  | 
(\<forall>X. equal(domain_of(restrct(element_relation::'a,universal_class,X)),sum_class(X))) &  | 
|
1409  | 
(\<forall>X. member(X::'a,universal_class) --> member(sum_class(X),universal_class)) &  | 
|
1410  | 
(\<forall>X. equal(complement(image'(element_relation::'a,complement(X))),powerClass(X))) &  | 
|
1411  | 
(\<forall>U. member(U::'a,universal_class) --> member(powerClass(U),universal_class)) &  | 
|
1412  | 
(\<forall>Yr Xr. subclass(compos(Yr::'a,Xr),cross_product(universal_class::'a,universal_class))) &  | 
|
1413  | 
(\<forall>Z Yr Xr Y. member(ordered_pair(Y::'a,Z),compos(Yr::'a,Xr)) --> member(Z::'a,image'(Yr::'a,image'(Xr::'a,singleton(Y))))) &  | 
|
1414  | 
(\<forall>Y Z Yr Xr. member(Z::'a,image'(Yr::'a,image'(Xr::'a,singleton(Y)))) & member(ordered_pair(Y::'a,Z),cross_product(universal_class::'a,universal_class)) --> member(ordered_pair(Y::'a,Z),compos(Yr::'a,Xr))) &  | 
|
1415  | 
(\<forall>X. single_valued_class(X) --> subclass(compos(X::'a,INVERSE(X)),identity_relation)) &  | 
|
1416  | 
(\<forall>X. subclass(compos(X::'a,INVERSE(X)),identity_relation) --> single_valued_class(X)) &  | 
|
1417  | 
(\<forall>Xf. function(Xf) --> subclass(Xf::'a,cross_product(universal_class::'a,universal_class))) &  | 
|
1418  | 
(\<forall>Xf. function(Xf) --> subclass(compos(Xf::'a,INVERSE(Xf)),identity_relation)) &  | 
|
1419  | 
(\<forall>Xf. subclass(Xf::'a,cross_product(universal_class::'a,universal_class)) & subclass(compos(Xf::'a,INVERSE(Xf)),identity_relation) --> function(Xf)) &  | 
|
1420  | 
(\<forall>Xf X. function(Xf) & member(X::'a,universal_class) --> member(image'(Xf::'a,X),universal_class)) &  | 
|
1421  | 
(\<forall>X. equal(X::'a,null_class) | member(regular(X),X)) &  | 
|
1422  | 
(\<forall>X. equal(X::'a,null_class) | equal(intersection(X::'a,regular(X)),null_class)) &  | 
|
1423  | 
(\<forall>Xf Y. equal(sum_class(image'(Xf::'a,singleton(Y))),apply(Xf::'a,Y))) &  | 
|
1424  | 
(function(choice)) &  | 
|
1425  | 
(\<forall>Y. member(Y::'a,universal_class) --> equal(Y::'a,null_class) | member(apply(choice::'a,Y),Y)) &  | 
|
1426  | 
(\<forall>Xf. one_to_one(Xf) --> function(Xf)) &  | 
|
1427  | 
(\<forall>Xf. one_to_one(Xf) --> function(INVERSE(Xf))) &  | 
|
1428  | 
(\<forall>Xf. function(INVERSE(Xf)) & function(Xf) --> one_to_one(Xf)) &  | 
|
1429  | 
(equal(intersection(cross_product(universal_class::'a,universal_class),intersection(cross_product(universal_class::'a,universal_class),complement(compos(complement(element_relation),INVERSE(element_relation))))),subset_relation)) &  | 
|
1430  | 
(equal(intersection(INVERSE(subset_relation),subset_relation),identity_relation)) &  | 
|
1431  | 
(\<forall>Xr. equal(complement(domain_of(intersection(Xr::'a,identity_relation))),diagonalise(Xr))) &  | 
|
1432  | 
(\<forall>X. equal(intersection(domain_of(X),diagonalise(compos(INVERSE(element_relation),X))),cantor(X))) &  | 
|
1433  | 
(\<forall>Xf. operation(Xf) --> function(Xf)) &  | 
|
1434  | 
(\<forall>Xf. operation(Xf) --> equal(cross_product(domain_of(domain_of(Xf)),domain_of(domain_of(Xf))),domain_of(Xf))) &  | 
|
1435  | 
(\<forall>Xf. operation(Xf) --> subclass(range_of(Xf),domain_of(domain_of(Xf)))) &  | 
|
1436  | 
(\<forall>Xf. function(Xf) & equal(cross_product(domain_of(domain_of(Xf)),domain_of(domain_of(Xf))),domain_of(Xf)) & subclass(range_of(Xf),domain_of(domain_of(Xf))) --> operation(Xf)) &  | 
|
1437  | 
(\<forall>Xf1 Xf2 Xh. compatible(Xh::'a,Xf1,Xf2) --> function(Xh)) &  | 
|
1438  | 
(\<forall>Xf2 Xf1 Xh. compatible(Xh::'a,Xf1,Xf2) --> equal(domain_of(domain_of(Xf1)),domain_of(Xh))) &  | 
|
1439  | 
(\<forall>Xf1 Xh Xf2. compatible(Xh::'a,Xf1,Xf2) --> subclass(range_of(Xh),domain_of(domain_of(Xf2)))) &  | 
|
1440  | 
(\<forall>Xh Xh1 Xf1 Xf2. function(Xh) & equal(domain_of(domain_of(Xf1)),domain_of(Xh)) & subclass(range_of(Xh),domain_of(domain_of(Xf2))) --> compatible(Xh1::'a,Xf1,Xf2)) &  | 
|
1441  | 
(\<forall>Xh Xf2 Xf1. homomorphism(Xh::'a,Xf1,Xf2) --> operation(Xf1)) &  | 
|
1442  | 
(\<forall>Xh Xf1 Xf2. homomorphism(Xh::'a,Xf1,Xf2) --> operation(Xf2)) &  | 
|
1443  | 
(\<forall>Xh Xf1 Xf2. homomorphism(Xh::'a,Xf1,Xf2) --> compatible(Xh::'a,Xf1,Xf2)) &  | 
|
1444  | 
(\<forall>Xf2 Xh Xf1 X Y. homomorphism(Xh::'a,Xf1,Xf2) & member(ordered_pair(X::'a,Y),domain_of(Xf1)) --> equal(apply(Xf2::'a,ordered_pair(apply(Xh::'a,X),apply(Xh::'a,Y))),apply(Xh::'a,apply(Xf1::'a,ordered_pair(X::'a,Y))))) &  | 
|
1445  | 
(\<forall>Xh Xf1 Xf2. operation(Xf1) & operation(Xf2) & compatible(Xh::'a,Xf1,Xf2) --> member(ordered_pair(not_homomorphism1(Xh::'a,Xf1,Xf2),not_homomorphism2(Xh::'a,Xf1,Xf2)),domain_of(Xf1)) | homomorphism(Xh::'a,Xf1,Xf2)) &  | 
|
| 24127 | 1446  | 
(\<forall>Xh Xf1 Xf2. operation(Xf1) & operation(Xf2) & compatible(Xh::'a,Xf1,Xf2) & equal(apply(Xf2::'a,ordered_pair(apply(Xh::'a,not_homomorphism1(Xh::'a,Xf1,Xf2)),apply(Xh::'a,not_homomorphism2(Xh::'a,Xf1,Xf2)))),apply(Xh::'a,apply(Xf1::'a,ordered_pair(not_homomorphism1(Xh::'a,Xf1,Xf2),not_homomorphism2(Xh::'a,Xf1,Xf2))))) --> homomorphism(Xh::'a,Xf1,Xf2))"  | 
1447  | 
||
1448  | 
abbreviation "SET004_0_eq subclass single_valued_class operation  | 
|
1449  | 
one_to_one member inductive homomorphism function compatible  | 
|
1450  | 
unordered_pair union sum_class successor singleton second rot restrct  | 
|
1451  | 
regular range_of rng powerClass ordered_pair not_subclass_element  | 
|
1452  | 
not_homomorphism2 not_homomorphism1 INVERSE intersection image' flip  | 
|
1453  | 
first domain_of domain difference diagonalise cross_product compos  | 
|
1454  | 
complement cantor apply equal \<equiv>  | 
|
| 24128 | 1455  | 
(\<forall>D E F'. equal(D::'a,E) --> equal(apply(D::'a,F'),apply(E::'a,F'))) &  | 
1456  | 
(\<forall>G I' H. equal(G::'a,H) --> equal(apply(I'::'a,G),apply(I'::'a,H))) &  | 
|
1457  | 
(\<forall>J K'. equal(J::'a,K') --> equal(cantor(J),cantor(K'))) &  | 
|
1458  | 
(\<forall>L M. equal(L::'a,M) --> equal(complement(L),complement(M))) &  | 
|
1459  | 
(\<forall>N O' P. equal(N::'a,O') --> equal(compos(N::'a,P),compos(O'::'a,P))) &  | 
|
1460  | 
(\<forall>Q S' R. equal(Q::'a,R) --> equal(compos(S'::'a,Q),compos(S'::'a,R))) &  | 
|
1461  | 
(\<forall>T' U V. equal(T'::'a,U) --> equal(cross_product(T'::'a,V),cross_product(U::'a,V))) &  | 
|
1462  | 
(\<forall>W Y X. equal(W::'a,X) --> equal(cross_product(Y::'a,W),cross_product(Y::'a,X))) &  | 
|
1463  | 
(\<forall>Z A1. equal(Z::'a,A1) --> equal(diagonalise(Z),diagonalise(A1))) &  | 
|
1464  | 
(\<forall>B1 C1 D1. equal(B1::'a,C1) --> equal(difference(B1::'a,D1),difference(C1::'a,D1))) &  | 
|
1465  | 
(\<forall>E1 G1 F1. equal(E1::'a,F1) --> equal(difference(G1::'a,E1),difference(G1::'a,F1))) &  | 
|
1466  | 
(\<forall>H1 I1 J1 K1. equal(H1::'a,I1) --> equal(domain(H1::'a,J1,K1),domain(I1::'a,J1,K1))) &  | 
|
1467  | 
(\<forall>L1 N1 M1 O1. equal(L1::'a,M1) --> equal(domain(N1::'a,L1,O1),domain(N1::'a,M1,O1))) &  | 
|
1468  | 
(\<forall>P1 R1 S1 Q1. equal(P1::'a,Q1) --> equal(domain(R1::'a,S1,P1),domain(R1::'a,S1,Q1))) &  | 
|
1469  | 
(\<forall>T1 U1. equal(T1::'a,U1) --> equal(domain_of(T1),domain_of(U1))) &  | 
|
1470  | 
(\<forall>V1 W1. equal(V1::'a,W1) --> equal(first(V1),first(W1))) &  | 
|
1471  | 
(\<forall>X1 Y1. equal(X1::'a,Y1) --> equal(flip(X1),flip(Y1))) &  | 
|
1472  | 
(\<forall>Z1 A2 B2. equal(Z1::'a,A2) --> equal(image'(Z1::'a,B2),image'(A2::'a,B2))) &  | 
|
1473  | 
(\<forall>C2 E2 D2. equal(C2::'a,D2) --> equal(image'(E2::'a,C2),image'(E2::'a,D2))) &  | 
|
1474  | 
(\<forall>F2 G2 H2. equal(F2::'a,G2) --> equal(intersection(F2::'a,H2),intersection(G2::'a,H2))) &  | 
|
1475  | 
(\<forall>I2 K2 J2. equal(I2::'a,J2) --> equal(intersection(K2::'a,I2),intersection(K2::'a,J2))) &  | 
|
1476  | 
(\<forall>L2 M2. equal(L2::'a,M2) --> equal(INVERSE(L2),INVERSE(M2))) &  | 
|
1477  | 
(\<forall>N2 O2 P2 Q2. equal(N2::'a,O2) --> equal(not_homomorphism1(N2::'a,P2,Q2),not_homomorphism1(O2::'a,P2,Q2))) &  | 
|
1478  | 
(\<forall>R2 T2 S2 U2. equal(R2::'a,S2) --> equal(not_homomorphism1(T2::'a,R2,U2),not_homomorphism1(T2::'a,S2,U2))) &  | 
|
1479  | 
(\<forall>V2 X2 Y2 W2. equal(V2::'a,W2) --> equal(not_homomorphism1(X2::'a,Y2,V2),not_homomorphism1(X2::'a,Y2,W2))) &  | 
|
1480  | 
(\<forall>Z2 A3 B3 C3. equal(Z2::'a,A3) --> equal(not_homomorphism2(Z2::'a,B3,C3),not_homomorphism2(A3::'a,B3,C3))) &  | 
|
1481  | 
(\<forall>D3 F3 E3 G3. equal(D3::'a,E3) --> equal(not_homomorphism2(F3::'a,D3,G3),not_homomorphism2(F3::'a,E3,G3))) &  | 
|
1482  | 
(\<forall>H3 J3 K3 I3. equal(H3::'a,I3) --> equal(not_homomorphism2(J3::'a,K3,H3),not_homomorphism2(J3::'a,K3,I3))) &  | 
|
1483  | 
(\<forall>L3 M3 N3. equal(L3::'a,M3) --> equal(not_subclass_element(L3::'a,N3),not_subclass_element(M3::'a,N3))) &  | 
|
1484  | 
(\<forall>O3 Q3 P3. equal(O3::'a,P3) --> equal(not_subclass_element(Q3::'a,O3),not_subclass_element(Q3::'a,P3))) &  | 
|
1485  | 
(\<forall>R3 S3 T3. equal(R3::'a,S3) --> equal(ordered_pair(R3::'a,T3),ordered_pair(S3::'a,T3))) &  | 
|
1486  | 
(\<forall>U3 W3 V3. equal(U3::'a,V3) --> equal(ordered_pair(W3::'a,U3),ordered_pair(W3::'a,V3))) &  | 
|
1487  | 
(\<forall>X3 Y3. equal(X3::'a,Y3) --> equal(powerClass(X3),powerClass(Y3))) &  | 
|
1488  | 
(\<forall>Z3 A4 B4 C4. equal(Z3::'a,A4) --> equal(rng(Z3::'a,B4,C4),rng(A4::'a,B4,C4))) &  | 
|
1489  | 
(\<forall>D4 F4 E4 G4. equal(D4::'a,E4) --> equal(rng(F4::'a,D4,G4),rng(F4::'a,E4,G4))) &  | 
|
1490  | 
(\<forall>H4 J4 K4 I4. equal(H4::'a,I4) --> equal(rng(J4::'a,K4,H4),rng(J4::'a,K4,I4))) &  | 
|
1491  | 
(\<forall>L4 M4. equal(L4::'a,M4) --> equal(range_of(L4),range_of(M4))) &  | 
|
1492  | 
(\<forall>N4 O4. equal(N4::'a,O4) --> equal(regular(N4),regular(O4))) &  | 
|
1493  | 
(\<forall>P4 Q4 R4 S4. equal(P4::'a,Q4) --> equal(restrct(P4::'a,R4,S4),restrct(Q4::'a,R4,S4))) &  | 
|
1494  | 
(\<forall>T4 V4 U4 W4. equal(T4::'a,U4) --> equal(restrct(V4::'a,T4,W4),restrct(V4::'a,U4,W4))) &  | 
|
1495  | 
(\<forall>X4 Z4 A5 Y4. equal(X4::'a,Y4) --> equal(restrct(Z4::'a,A5,X4),restrct(Z4::'a,A5,Y4))) &  | 
|
1496  | 
(\<forall>B5 C5. equal(B5::'a,C5) --> equal(rot(B5),rot(C5))) &  | 
|
1497  | 
(\<forall>D5 E5. equal(D5::'a,E5) --> equal(second(D5),second(E5))) &  | 
|
1498  | 
(\<forall>F5 G5. equal(F5::'a,G5) --> equal(singleton(F5),singleton(G5))) &  | 
|
1499  | 
(\<forall>H5 I5. equal(H5::'a,I5) --> equal(successor(H5),successor(I5))) &  | 
|
1500  | 
(\<forall>J5 K5. equal(J5::'a,K5) --> equal(sum_class(J5),sum_class(K5))) &  | 
|
1501  | 
(\<forall>L5 M5 N5. equal(L5::'a,M5) --> equal(union(L5::'a,N5),union(M5::'a,N5))) &  | 
|
1502  | 
(\<forall>O5 Q5 P5. equal(O5::'a,P5) --> equal(union(Q5::'a,O5),union(Q5::'a,P5))) &  | 
|
1503  | 
(\<forall>R5 S5 T5. equal(R5::'a,S5) --> equal(unordered_pair(R5::'a,T5),unordered_pair(S5::'a,T5))) &  | 
|
1504  | 
(\<forall>U5 W5 V5. equal(U5::'a,V5) --> equal(unordered_pair(W5::'a,U5),unordered_pair(W5::'a,V5))) &  | 
|
1505  | 
(\<forall>X5 Y5 Z5 A6. equal(X5::'a,Y5) & compatible(X5::'a,Z5,A6) --> compatible(Y5::'a,Z5,A6)) &  | 
|
1506  | 
(\<forall>B6 D6 C6 E6. equal(B6::'a,C6) & compatible(D6::'a,B6,E6) --> compatible(D6::'a,C6,E6)) &  | 
|
1507  | 
(\<forall>F6 H6 I6 G6. equal(F6::'a,G6) & compatible(H6::'a,I6,F6) --> compatible(H6::'a,I6,G6)) &  | 
|
1508  | 
(\<forall>J6 K6. equal(J6::'a,K6) & function(J6) --> function(K6)) &  | 
|
1509  | 
(\<forall>L6 M6 N6 O6. equal(L6::'a,M6) & homomorphism(L6::'a,N6,O6) --> homomorphism(M6::'a,N6,O6)) &  | 
|
1510  | 
(\<forall>P6 R6 Q6 S6. equal(P6::'a,Q6) & homomorphism(R6::'a,P6,S6) --> homomorphism(R6::'a,Q6,S6)) &  | 
|
1511  | 
(\<forall>T6 V6 W6 U6. equal(T6::'a,U6) & homomorphism(V6::'a,W6,T6) --> homomorphism(V6::'a,W6,U6)) &  | 
|
1512  | 
(\<forall>X6 Y6. equal(X6::'a,Y6) & inductive(X6) --> inductive(Y6)) &  | 
|
1513  | 
(\<forall>Z6 A7 B7. equal(Z6::'a,A7) & member(Z6::'a,B7) --> member(A7::'a,B7)) &  | 
|
1514  | 
(\<forall>C7 E7 D7. equal(C7::'a,D7) & member(E7::'a,C7) --> member(E7::'a,D7)) &  | 
|
1515  | 
(\<forall>F7 G7. equal(F7::'a,G7) & one_to_one(F7) --> one_to_one(G7)) &  | 
|
1516  | 
(\<forall>H7 I7. equal(H7::'a,I7) & operation(H7) --> operation(I7)) &  | 
|
1517  | 
(\<forall>J7 K7. equal(J7::'a,K7) & single_valued_class(J7) --> single_valued_class(K7)) &  | 
|
1518  | 
(\<forall>L7 M7 N7. equal(L7::'a,M7) & subclass(L7::'a,N7) --> subclass(M7::'a,N7)) &  | 
|
1519  | 
(\<forall>O7 Q7 P7. equal(O7::'a,P7) & subclass(Q7::'a,O7) --> subclass(Q7::'a,P7))"  | 
|
| 24127 | 1520  | 
|
1521  | 
abbreviation "SET004_1_ax range_of function maps apply  | 
|
1522  | 
application_function singleton_relation element_relation complement  | 
|
1523  | 
intersection single_valued3 singleton image' domain single_valued2  | 
|
1524  | 
second single_valued1 identity_relation INVERSE not_subclass_element  | 
|
1525  | 
first domain_of domain_relation composition_function compos equal  | 
|
1526  | 
ordered_pair member universal_class cross_product compose_class  | 
|
1527  | 
subclass \<equiv>  | 
|
| 24128 | 1528  | 
(\<forall>X. subclass(compose_class(X),cross_product(universal_class::'a,universal_class))) &  | 
1529  | 
(\<forall>X Y Z. member(ordered_pair(Y::'a,Z),compose_class(X)) --> equal(compos(X::'a,Y),Z)) &  | 
|
1530  | 
(\<forall>Y Z X. member(ordered_pair(Y::'a,Z),cross_product(universal_class::'a,universal_class)) & equal(compos(X::'a,Y),Z) --> member(ordered_pair(Y::'a,Z),compose_class(X))) &  | 
|
1531  | 
(subclass(composition_function::'a,cross_product(universal_class::'a,cross_product(universal_class::'a,universal_class)))) &  | 
|
1532  | 
(\<forall>X Y Z. member(ordered_pair(X::'a,ordered_pair(Y::'a,Z)),composition_function) --> equal(compos(X::'a,Y),Z)) &  | 
|
1533  | 
(\<forall>X Y. member(ordered_pair(X::'a,Y),cross_product(universal_class::'a,universal_class)) --> member(ordered_pair(X::'a,ordered_pair(Y::'a,compos(X::'a,Y))),composition_function)) &  | 
|
1534  | 
(subclass(domain_relation::'a,cross_product(universal_class::'a,universal_class))) &  | 
|
1535  | 
(\<forall>X Y. member(ordered_pair(X::'a,Y),domain_relation) --> equal(domain_of(X),Y)) &  | 
|
1536  | 
(\<forall>X. member(X::'a,universal_class) --> member(ordered_pair(X::'a,domain_of(X)),domain_relation)) &  | 
|
1537  | 
(\<forall>X. equal(first(not_subclass_element(compos(X::'a,INVERSE(X)),identity_relation)),single_valued1(X))) &  | 
|
1538  | 
(\<forall>X. equal(second(not_subclass_element(compos(X::'a,INVERSE(X)),identity_relation)),single_valued2(X))) &  | 
|
1539  | 
(\<forall>X. equal(domain(X::'a,image'(INVERSE(X),singleton(single_valued1(X))),single_valued2(X)),single_valued3(X))) &  | 
|
1540  | 
(equal(intersection(complement(compos(element_relation::'a,complement(identity_relation))),element_relation),singleton_relation)) &  | 
|
1541  | 
(subclass(application_function::'a,cross_product(universal_class::'a,cross_product(universal_class::'a,universal_class)))) &  | 
|
1542  | 
(\<forall>Z Y X. member(ordered_pair(X::'a,ordered_pair(Y::'a,Z)),application_function) --> member(Y::'a,domain_of(X))) &  | 
|
1543  | 
(\<forall>X Y Z. member(ordered_pair(X::'a,ordered_pair(Y::'a,Z)),application_function) --> equal(apply(X::'a,Y),Z)) &  | 
|
1544  | 
(\<forall>Z X Y. member(ordered_pair(X::'a,ordered_pair(Y::'a,Z)),cross_product(universal_class::'a,cross_product(universal_class::'a,universal_class))) & member(Y::'a,domain_of(X)) --> member(ordered_pair(X::'a,ordered_pair(Y::'a,apply(X::'a,Y))),application_function)) &  | 
|
1545  | 
(\<forall>X Y Xf. maps(Xf::'a,X,Y) --> function(Xf)) &  | 
|
1546  | 
(\<forall>Y Xf X. maps(Xf::'a,X,Y) --> equal(domain_of(Xf),X)) &  | 
|
1547  | 
(\<forall>X Xf Y. maps(Xf::'a,X,Y) --> subclass(range_of(Xf),Y)) &  | 
|
| 24127 | 1548  | 
(\<forall>Xf Y. function(Xf) & subclass(range_of(Xf),Y) --> maps(Xf::'a,domain_of(Xf),Y))"  | 
1549  | 
||
1550  | 
abbreviation "SET004_1_eq maps single_valued3 single_valued2 single_valued1 compose_class equal \<equiv>  | 
|
| 24128 | 1551  | 
(\<forall>L M. equal(L::'a,M) --> equal(compose_class(L),compose_class(M))) &  | 
1552  | 
(\<forall>N2 O2. equal(N2::'a,O2) --> equal(single_valued1(N2),single_valued1(O2))) &  | 
|
1553  | 
(\<forall>P2 Q2. equal(P2::'a,Q2) --> equal(single_valued2(P2),single_valued2(Q2))) &  | 
|
1554  | 
(\<forall>R2 S2. equal(R2::'a,S2) --> equal(single_valued3(R2),single_valued3(S2))) &  | 
|
1555  | 
(\<forall>X2 Y2 Z2 A3. equal(X2::'a,Y2) & maps(X2::'a,Z2,A3) --> maps(Y2::'a,Z2,A3)) &  | 
|
1556  | 
(\<forall>B3 D3 C3 E3. equal(B3::'a,C3) & maps(D3::'a,B3,E3) --> maps(D3::'a,C3,E3)) &  | 
|
| 24127 | 1557  | 
(\<forall>F3 H3 I3 G3. equal(F3::'a,G3) & maps(H3::'a,I3,F3) --> maps(H3::'a,I3,G3))"  | 
1558  | 
||
1559  | 
abbreviation "NUM004_0_ax integer_of omega ordinal_multiply  | 
|
1560  | 
add_relation ordinal_add recursion apply range_of union_of_range_map  | 
|
1561  | 
function recursion_equation_functions rest_relation rest_of  | 
|
1562  | 
limit_ordinals kind_1_ordinals successor_relation image'  | 
|
1563  | 
universal_class sum_class element_relation ordinal_numbers section  | 
|
1564  | 
not_well_ordering ordered_pair least member well_ordering singleton  | 
|
1565  | 
domain_of segment null_class intersection asymmetric compos transitive  | 
|
1566  | 
cross_product connected identity_relation complement restrct subclass  | 
|
1567  | 
irreflexive symmetrization_of INVERSE union equal \<equiv>  | 
|
| 24128 | 1568  | 
(\<forall>X. equal(union(X::'a,INVERSE(X)),symmetrization_of(X))) &  | 
1569  | 
(\<forall>X Y. irreflexive(X::'a,Y) --> subclass(restrct(X::'a,Y,Y),complement(identity_relation))) &  | 
|
1570  | 
(\<forall>X Y. subclass(restrct(X::'a,Y,Y),complement(identity_relation)) --> irreflexive(X::'a,Y)) &  | 
|
1571  | 
(\<forall>Y X. connected(X::'a,Y) --> subclass(cross_product(Y::'a,Y),union(identity_relation::'a,symmetrization_of(X)))) &  | 
|
1572  | 
(\<forall>X Y. subclass(cross_product(Y::'a,Y),union(identity_relation::'a,symmetrization_of(X))) --> connected(X::'a,Y)) &  | 
|
1573  | 
(\<forall>Xr Y. transitive(Xr::'a,Y) --> subclass(compos(restrct(Xr::'a,Y,Y),restrct(Xr::'a,Y,Y)),restrct(Xr::'a,Y,Y))) &  | 
|
1574  | 
(\<forall>Xr Y. subclass(compos(restrct(Xr::'a,Y,Y),restrct(Xr::'a,Y,Y)),restrct(Xr::'a,Y,Y)) --> transitive(Xr::'a,Y)) &  | 
|
1575  | 
(\<forall>Xr Y. asymmetric(Xr::'a,Y) --> equal(restrct(intersection(Xr::'a,INVERSE(Xr)),Y,Y),null_class)) &  | 
|
1576  | 
(\<forall>Xr Y. equal(restrct(intersection(Xr::'a,INVERSE(Xr)),Y,Y),null_class) --> asymmetric(Xr::'a,Y)) &  | 
|
1577  | 
(\<forall>Xr Y Z. equal(segment(Xr::'a,Y,Z),domain_of(restrct(Xr::'a,Y,singleton(Z))))) &  | 
|
1578  | 
(\<forall>X Y. well_ordering(X::'a,Y) --> connected(X::'a,Y)) &  | 
|
1579  | 
(\<forall>Y Xr U. well_ordering(Xr::'a,Y) & subclass(U::'a,Y) --> equal(U::'a,null_class) | member(least(Xr::'a,U),U)) &  | 
|
1580  | 
(\<forall>Y V Xr U. well_ordering(Xr::'a,Y) & subclass(U::'a,Y) & member(V::'a,U) --> member(least(Xr::'a,U),U)) &  | 
|
1581  | 
(\<forall>Y Xr U. well_ordering(Xr::'a,Y) & subclass(U::'a,Y) --> equal(segment(Xr::'a,U,least(Xr::'a,U)),null_class)) &  | 
|
1582  | 
(\<forall>Y V U Xr. ~(well_ordering(Xr::'a,Y) & subclass(U::'a,Y) & member(V::'a,U) & member(ordered_pair(V::'a,least(Xr::'a,U)),Xr))) &  | 
|
1583  | 
(\<forall>Xr Y. connected(Xr::'a,Y) & equal(not_well_ordering(Xr::'a,Y),null_class) --> well_ordering(Xr::'a,Y)) &  | 
|
1584  | 
(\<forall>Xr Y. connected(Xr::'a,Y) --> subclass(not_well_ordering(Xr::'a,Y),Y) | well_ordering(Xr::'a,Y)) &  | 
|
1585  | 
(\<forall>V Xr Y. member(V::'a,not_well_ordering(Xr::'a,Y)) & equal(segment(Xr::'a,not_well_ordering(Xr::'a,Y),V),null_class) & connected(Xr::'a,Y) --> well_ordering(Xr::'a,Y)) &  | 
|
1586  | 
(\<forall>Xr Y Z. section(Xr::'a,Y,Z) --> subclass(Y::'a,Z)) &  | 
|
1587  | 
(\<forall>Xr Z Y. section(Xr::'a,Y,Z) --> subclass(domain_of(restrct(Xr::'a,Z,Y)),Y)) &  | 
|
1588  | 
(\<forall>Xr Y Z. subclass(Y::'a,Z) & subclass(domain_of(restrct(Xr::'a,Z,Y)),Y) --> section(Xr::'a,Y,Z)) &  | 
|
1589  | 
(\<forall>X. member(X::'a,ordinal_numbers) --> well_ordering(element_relation::'a,X)) &  | 
|
1590  | 
(\<forall>X. member(X::'a,ordinal_numbers) --> subclass(sum_class(X),X)) &  | 
|
1591  | 
(\<forall>X. well_ordering(element_relation::'a,X) & subclass(sum_class(X),X) & member(X::'a,universal_class) --> member(X::'a,ordinal_numbers)) &  | 
|
1592  | 
(\<forall>X. well_ordering(element_relation::'a,X) & subclass(sum_class(X),X) --> member(X::'a,ordinal_numbers) | equal(X::'a,ordinal_numbers)) &  | 
|
1593  | 
(equal(union(singleton(null_class),image'(successor_relation::'a,ordinal_numbers)),kind_1_ordinals)) &  | 
|
1594  | 
(equal(intersection(complement(kind_1_ordinals),ordinal_numbers),limit_ordinals)) &  | 
|
1595  | 
(\<forall>X. subclass(rest_of(X),cross_product(universal_class::'a,universal_class))) &  | 
|
1596  | 
(\<forall>V U X. member(ordered_pair(U::'a,V),rest_of(X)) --> member(U::'a,domain_of(X))) &  | 
|
1597  | 
(\<forall>X U V. member(ordered_pair(U::'a,V),rest_of(X)) --> equal(restrct(X::'a,U,universal_class),V)) &  | 
|
1598  | 
(\<forall>U V X. member(U::'a,domain_of(X)) & equal(restrct(X::'a,U,universal_class),V) --> member(ordered_pair(U::'a,V),rest_of(X))) &  | 
|
1599  | 
(subclass(rest_relation::'a,cross_product(universal_class::'a,universal_class))) &  | 
|
1600  | 
(\<forall>X Y. member(ordered_pair(X::'a,Y),rest_relation) --> equal(rest_of(X),Y)) &  | 
|
1601  | 
(\<forall>X. member(X::'a,universal_class) --> member(ordered_pair(X::'a,rest_of(X)),rest_relation)) &  | 
|
1602  | 
(\<forall>X Z. member(X::'a,recursion_equation_functions(Z)) --> function(Z)) &  | 
|
1603  | 
(\<forall>Z X. member(X::'a,recursion_equation_functions(Z)) --> function(X)) &  | 
|
1604  | 
(\<forall>Z X. member(X::'a,recursion_equation_functions(Z)) --> member(domain_of(X),ordinal_numbers)) &  | 
|
1605  | 
(\<forall>Z X. member(X::'a,recursion_equation_functions(Z)) --> equal(compos(Z::'a,rest_of(X)),X)) &  | 
|
1606  | 
(\<forall>X Z. function(Z) & function(X) & member(domain_of(X),ordinal_numbers) & equal(compos(Z::'a,rest_of(X)),X) --> member(X::'a,recursion_equation_functions(Z))) &  | 
|
1607  | 
(subclass(union_of_range_map::'a,cross_product(universal_class::'a,universal_class))) &  | 
|
1608  | 
(\<forall>X Y. member(ordered_pair(X::'a,Y),union_of_range_map) --> equal(sum_class(range_of(X)),Y)) &  | 
|
1609  | 
(\<forall>X Y. member(ordered_pair(X::'a,Y),cross_product(universal_class::'a,universal_class)) & equal(sum_class(range_of(X)),Y) --> member(ordered_pair(X::'a,Y),union_of_range_map)) &  | 
|
1610  | 
(\<forall>X Y. equal(apply(recursion(X::'a,successor_relation,union_of_range_map),Y),ordinal_add(X::'a,Y))) &  | 
|
1611  | 
(\<forall>X Y. equal(recursion(null_class::'a,apply(add_relation::'a,X),union_of_range_map),ordinal_multiply(X::'a,Y))) &  | 
|
1612  | 
(\<forall>X. member(X::'a,omega) --> equal(integer_of(X),X)) &  | 
|
| 24127 | 1613  | 
(\<forall>X. member(X::'a,omega) | equal(integer_of(X),null_class))"  | 
1614  | 
||
1615  | 
abbreviation "NUM004_0_eq well_ordering transitive section irreflexive  | 
|
1616  | 
connected asymmetric symmetrization_of segment rest_of  | 
|
1617  | 
recursion_equation_functions recursion ordinal_multiply ordinal_add  | 
|
1618  | 
not_well_ordering least integer_of equal \<equiv>  | 
|
| 24128 | 1619  | 
(\<forall>D E. equal(D::'a,E) --> equal(integer_of(D),integer_of(E))) &  | 
1620  | 
(\<forall>F' G H. equal(F'::'a,G) --> equal(least(F'::'a,H),least(G::'a,H))) &  | 
|
1621  | 
(\<forall>I' K' J. equal(I'::'a,J) --> equal(least(K'::'a,I'),least(K'::'a,J))) &  | 
|
1622  | 
(\<forall>L M N. equal(L::'a,M) --> equal(not_well_ordering(L::'a,N),not_well_ordering(M::'a,N))) &  | 
|
1623  | 
(\<forall>O' Q P. equal(O'::'a,P) --> equal(not_well_ordering(Q::'a,O'),not_well_ordering(Q::'a,P))) &  | 
|
1624  | 
(\<forall>R S' T'. equal(R::'a,S') --> equal(ordinal_add(R::'a,T'),ordinal_add(S'::'a,T'))) &  | 
|
1625  | 
(\<forall>U W V. equal(U::'a,V) --> equal(ordinal_add(W::'a,U),ordinal_add(W::'a,V))) &  | 
|
1626  | 
(\<forall>X Y Z. equal(X::'a,Y) --> equal(ordinal_multiply(X::'a,Z),ordinal_multiply(Y::'a,Z))) &  | 
|
1627  | 
(\<forall>A1 C1 B1. equal(A1::'a,B1) --> equal(ordinal_multiply(C1::'a,A1),ordinal_multiply(C1::'a,B1))) &  | 
|
1628  | 
(\<forall>F1 G1 H1 I1. equal(F1::'a,G1) --> equal(recursion(F1::'a,H1,I1),recursion(G1::'a,H1,I1))) &  | 
|
1629  | 
(\<forall>J1 L1 K1 M1. equal(J1::'a,K1) --> equal(recursion(L1::'a,J1,M1),recursion(L1::'a,K1,M1))) &  | 
|
1630  | 
(\<forall>N1 P1 Q1 O1. equal(N1::'a,O1) --> equal(recursion(P1::'a,Q1,N1),recursion(P1::'a,Q1,O1))) &  | 
|
1631  | 
(\<forall>R1 S1. equal(R1::'a,S1) --> equal(recursion_equation_functions(R1),recursion_equation_functions(S1))) &  | 
|
1632  | 
(\<forall>T1 U1. equal(T1::'a,U1) --> equal(rest_of(T1),rest_of(U1))) &  | 
|
1633  | 
(\<forall>V1 W1 X1 Y1. equal(V1::'a,W1) --> equal(segment(V1::'a,X1,Y1),segment(W1::'a,X1,Y1))) &  | 
|
1634  | 
(\<forall>Z1 B2 A2 C2. equal(Z1::'a,A2) --> equal(segment(B2::'a,Z1,C2),segment(B2::'a,A2,C2))) &  | 
|
1635  | 
(\<forall>D2 F2 G2 E2. equal(D2::'a,E2) --> equal(segment(F2::'a,G2,D2),segment(F2::'a,G2,E2))) &  | 
|
1636  | 
(\<forall>H2 I2. equal(H2::'a,I2) --> equal(symmetrization_of(H2),symmetrization_of(I2))) &  | 
|
1637  | 
(\<forall>J2 K2 L2. equal(J2::'a,K2) & asymmetric(J2::'a,L2) --> asymmetric(K2::'a,L2)) &  | 
|
1638  | 
(\<forall>M2 O2 N2. equal(M2::'a,N2) & asymmetric(O2::'a,M2) --> asymmetric(O2::'a,N2)) &  | 
|
1639  | 
(\<forall>P2 Q2 R2. equal(P2::'a,Q2) & connected(P2::'a,R2) --> connected(Q2::'a,R2)) &  | 
|
1640  | 
(\<forall>S2 U2 T2. equal(S2::'a,T2) & connected(U2::'a,S2) --> connected(U2::'a,T2)) &  | 
|
1641  | 
(\<forall>V2 W2 X2. equal(V2::'a,W2) & irreflexive(V2::'a,X2) --> irreflexive(W2::'a,X2)) &  | 
|
1642  | 
(\<forall>Y2 A3 Z2. equal(Y2::'a,Z2) & irreflexive(A3::'a,Y2) --> irreflexive(A3::'a,Z2)) &  | 
|
1643  | 
(\<forall>B3 C3 D3 E3. equal(B3::'a,C3) & section(B3::'a,D3,E3) --> section(C3::'a,D3,E3)) &  | 
|
1644  | 
(\<forall>F3 H3 G3 I3. equal(F3::'a,G3) & section(H3::'a,F3,I3) --> section(H3::'a,G3,I3)) &  | 
|
1645  | 
(\<forall>J3 L3 M3 K3. equal(J3::'a,K3) & section(L3::'a,M3,J3) --> section(L3::'a,M3,K3)) &  | 
|
1646  | 
(\<forall>N3 O3 P3. equal(N3::'a,O3) & transitive(N3::'a,P3) --> transitive(O3::'a,P3)) &  | 
|
1647  | 
(\<forall>Q3 S3 R3. equal(Q3::'a,R3) & transitive(S3::'a,Q3) --> transitive(S3::'a,R3)) &  | 
|
1648  | 
(\<forall>T3 U3 V3. equal(T3::'a,U3) & well_ordering(T3::'a,V3) --> well_ordering(U3::'a,V3)) &  | 
|
| 24127 | 1649  | 
(\<forall>W3 Y3 X3. equal(W3::'a,X3) & well_ordering(Y3::'a,W3) --> well_ordering(Y3::'a,X3))"  | 
1650  | 
||
1651  | 
(*1345 inferences so far. Searching to depth 7. 23.3 secs. BIG*)  | 
|
1652  | 
lemma NUM180_1:  | 
|
1653  | 
"EQU001_0_ax equal &  | 
|
1654  | 
SET004_0_ax not_homomorphism2 not_homomorphism1  | 
|
1655  | 
homomorphism compatible operation cantor diagonalise subset_relation  | 
|
1656  | 
one_to_one choice apply regular function identity_relation  | 
|
1657  | 
single_valued_class compos powerClass sum_class omega inductive  | 
|
1658  | 
successor_relation successor image' rng domain range_of INVERSE flip  | 
|
1659  | 
rot domain_of null_class restrct difference union complement  | 
|
1660  | 
intersection element_relation second first cross_product ordered_pair  | 
|
1661  | 
singleton unordered_pair equal universal_class not_subclass_element  | 
|
1662  | 
member subclass &  | 
|
1663  | 
SET004_0_eq subclass single_valued_class operation  | 
|
1664  | 
one_to_one member inductive homomorphism function compatible  | 
|
1665  | 
unordered_pair union sum_class successor singleton second rot restrct  | 
|
1666  | 
regular range_of rng powerClass ordered_pair not_subclass_element  | 
|
1667  | 
not_homomorphism2 not_homomorphism1 INVERSE intersection image' flip  | 
|
1668  | 
first domain_of domain difference diagonalise cross_product compos  | 
|
1669  | 
complement cantor apply equal &  | 
|
1670  | 
SET004_1_ax range_of function maps apply  | 
|
1671  | 
application_function singleton_relation element_relation complement  | 
|
1672  | 
intersection single_valued3 singleton image' domain single_valued2  | 
|
1673  | 
second single_valued1 identity_relation INVERSE not_subclass_element  | 
|
1674  | 
first domain_of domain_relation composition_function compos equal  | 
|
1675  | 
ordered_pair member universal_class cross_product compose_class  | 
|
1676  | 
subclass &  | 
|
1677  | 
SET004_1_eq maps single_valued3 single_valued2 single_valued1 compose_class equal &  | 
|
1678  | 
NUM004_0_ax integer_of omega ordinal_multiply  | 
|
1679  | 
add_relation ordinal_add recursion apply range_of union_of_range_map  | 
|
1680  | 
function recursion_equation_functions rest_relation rest_of  | 
|
1681  | 
limit_ordinals kind_1_ordinals successor_relation image'  | 
|
1682  | 
universal_class sum_class element_relation ordinal_numbers section  | 
|
1683  | 
not_well_ordering ordered_pair least member well_ordering singleton  | 
|
1684  | 
domain_of segment null_class intersection asymmetric compos transitive  | 
|
1685  | 
cross_product connected identity_relation complement restrct subclass  | 
|
1686  | 
irreflexive symmetrization_of INVERSE union equal &  | 
|
1687  | 
NUM004_0_eq well_ordering transitive section irreflexive  | 
|
1688  | 
connected asymmetric symmetrization_of segment rest_of  | 
|
1689  | 
recursion_equation_functions recursion ordinal_multiply ordinal_add  | 
|
1690  | 
not_well_ordering least integer_of equal &  | 
|
1691  | 
(~subclass(limit_ordinals::'a,ordinal_numbers)) --> False"  | 
|
1692  | 
by meson  | 
|
1693  | 
||
1694  | 
||
1695  | 
(*0 inferences so far. Searching to depth 0. 16.8 secs. BIG*)  | 
|
1696  | 
lemma NUM228_1:  | 
|
1697  | 
"EQU001_0_ax equal &  | 
|
1698  | 
SET004_0_ax not_homomorphism2 not_homomorphism1  | 
|
1699  | 
homomorphism compatible operation cantor diagonalise subset_relation  | 
|
1700  | 
one_to_one choice apply regular function identity_relation  | 
|
1701  | 
single_valued_class compos powerClass sum_class omega inductive  | 
|
1702  | 
successor_relation successor image' rng domain range_of INVERSE flip  | 
|
1703  | 
rot domain_of null_class restrct difference union complement  | 
|
1704  | 
intersection element_relation second first cross_product ordered_pair  | 
|
1705  | 
singleton unordered_pair equal universal_class not_subclass_element  | 
|
1706  | 
member subclass &  | 
|
1707  | 
SET004_0_eq subclass single_valued_class operation  | 
|
1708  | 
one_to_one member inductive homomorphism function compatible  | 
|
1709  | 
unordered_pair union sum_class successor singleton second rot restrct  | 
|
1710  | 
regular range_of rng powerClass ordered_pair not_subclass_element  | 
|
1711  | 
not_homomorphism2 not_homomorphism1 INVERSE intersection image' flip  | 
|
1712  | 
first domain_of domain difference diagonalise cross_product compos  | 
|
1713  | 
complement cantor apply equal &  | 
|
1714  | 
SET004_1_ax range_of function maps apply  | 
|
1715  | 
application_function singleton_relation element_relation complement  | 
|
1716  | 
intersection single_valued3 singleton image' domain single_valued2  | 
|
1717  | 
second single_valued1 identity_relation INVERSE not_subclass_element  | 
|
1718  | 
first domain_of domain_relation composition_function compos equal  | 
|
1719  | 
ordered_pair member universal_class cross_product compose_class  | 
|
1720  | 
subclass &  | 
|
1721  | 
SET004_1_eq maps single_valued3 single_valued2 single_valued1 compose_class equal &  | 
|
1722  | 
NUM004_0_ax integer_of omega ordinal_multiply  | 
|
1723  | 
add_relation ordinal_add recursion apply range_of union_of_range_map  | 
|
1724  | 
function recursion_equation_functions rest_relation rest_of  | 
|
1725  | 
limit_ordinals kind_1_ordinals successor_relation image'  | 
|
1726  | 
universal_class sum_class element_relation ordinal_numbers section  | 
|
1727  | 
not_well_ordering ordered_pair least member well_ordering singleton  | 
|
1728  | 
domain_of segment null_class intersection asymmetric compos transitive  | 
|
1729  | 
cross_product connected identity_relation complement restrct subclass  | 
|
1730  | 
irreflexive symmetrization_of INVERSE union equal &  | 
|
1731  | 
NUM004_0_eq well_ordering transitive section irreflexive  | 
|
1732  | 
connected asymmetric symmetrization_of segment rest_of  | 
|
1733  | 
recursion_equation_functions recursion ordinal_multiply ordinal_add  | 
|
1734  | 
not_well_ordering least integer_of equal &  | 
|
| 24128 | 1735  | 
(~function(z)) &  | 
| 24127 | 1736  | 
(~equal(recursion_equation_functions(z),null_class)) --> False"  | 
1737  | 
by meson  | 
|
1738  | 
||
1739  | 
||
1740  | 
(*4868 inferences so far. Searching to depth 12. 4.3 secs*)  | 
|
1741  | 
lemma PLA002_1:  | 
|
| 24128 | 1742  | 
"(\<forall>Situation1 Situation2. warm(Situation1) | cold(Situation2)) &  | 
1743  | 
(\<forall>Situation. at(a::'a,Situation) --> at(b::'a,walk(b::'a,Situation))) &  | 
|
1744  | 
(\<forall>Situation. at(a::'a,Situation) --> at(b::'a,drive(b::'a,Situation))) &  | 
|
1745  | 
(\<forall>Situation. at(b::'a,Situation) --> at(a::'a,walk(a::'a,Situation))) &  | 
|
1746  | 
(\<forall>Situation. at(b::'a,Situation) --> at(a::'a,drive(a::'a,Situation))) &  | 
|
1747  | 
(\<forall>Situation. cold(Situation) & at(b::'a,Situation) --> at(c::'a,skate(c::'a,Situation))) &  | 
|
1748  | 
(\<forall>Situation. cold(Situation) & at(c::'a,Situation) --> at(b::'a,skate(b::'a,Situation))) &  | 
|
1749  | 
(\<forall>Situation. warm(Situation) & at(b::'a,Situation) --> at(d::'a,climb(d::'a,Situation))) &  | 
|
1750  | 
(\<forall>Situation. warm(Situation) & at(d::'a,Situation) --> at(b::'a,climb(b::'a,Situation))) &  | 
|
1751  | 
(\<forall>Situation. at(c::'a,Situation) --> at(d::'a,go(d::'a,Situation))) &  | 
|
1752  | 
(\<forall>Situation. at(d::'a,Situation) --> at(c::'a,go(c::'a,Situation))) &  | 
|
1753  | 
(\<forall>Situation. at(c::'a,Situation) --> at(e::'a,go(e::'a,Situation))) &  | 
|
1754  | 
(\<forall>Situation. at(e::'a,Situation) --> at(c::'a,go(c::'a,Situation))) &  | 
|
1755  | 
(\<forall>Situation. at(d::'a,Situation) --> at(f::'a,go(f::'a,Situation))) &  | 
|
1756  | 
(\<forall>Situation. at(f::'a,Situation) --> at(d::'a,go(d::'a,Situation))) &  | 
|
1757  | 
(at(f::'a,s0)) &  | 
|
| 24127 | 1758  | 
(\<forall>S'. ~at(a::'a,S')) --> False"  | 
1759  | 
by meson  | 
|
1760  | 
||
1761  | 
abbreviation "PLA001_0_ax putdown on pickup do holding table differ clear EMPTY and' holds \<equiv>  | 
|
| 24128 | 1762  | 
(\<forall>X Y State. holds(X::'a,State) & holds(Y::'a,State) --> holds(and'(X::'a,Y),State)) &  | 
1763  | 
(\<forall>State X. holds(EMPTY::'a,State) & holds(clear(X),State) & differ(X::'a,table) --> holds(holding(X),do(pickup(X),State))) &  | 
|
1764  | 
(\<forall>Y X State. holds(on(X::'a,Y),State) & holds(clear(X),State) & holds(EMPTY::'a,State) --> holds(clear(Y),do(pickup(X),State))) &  | 
|
1765  | 
(\<forall>Y State X Z. holds(on(X::'a,Y),State) & differ(X::'a,Z) --> holds(on(X::'a,Y),do(pickup(Z),State))) &  | 
|
1766  | 
(\<forall>State X Z. holds(clear(X),State) & differ(X::'a,Z) --> holds(clear(X),do(pickup(Z),State))) &  | 
|
1767  | 
(\<forall>X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(EMPTY::'a,do(putdown(X::'a,Y),State))) &  | 
|
1768  | 
(\<forall>X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(on(X::'a,Y),do(putdown(X::'a,Y),State))) &  | 
|
1769  | 
(\<forall>X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(clear(X),do(putdown(X::'a,Y),State))) &  | 
|
1770  | 
(\<forall>Z W X Y State. holds(on(X::'a,Y),State) --> holds(on(X::'a,Y),do(putdown(Z::'a,W),State))) &  | 
|
| 24127 | 1771  | 
(\<forall>X State Z Y. holds(clear(Z),State) & differ(Z::'a,Y) --> holds(clear(Z),do(putdown(X::'a,Y),State)))"  | 
1772  | 
||
1773  | 
abbreviation "PLA001_1_ax EMPTY clear s0 on holds table d c b a differ \<equiv>  | 
|
| 24128 | 1774  | 
(\<forall>Y X. differ(Y::'a,X) --> differ(X::'a,Y)) &  | 
1775  | 
(differ(a::'a,b)) &  | 
|
1776  | 
(differ(a::'a,c)) &  | 
|
1777  | 
(differ(a::'a,d)) &  | 
|
1778  | 
(differ(a::'a,table)) &  | 
|
1779  | 
(differ(b::'a,c)) &  | 
|
1780  | 
(differ(b::'a,d)) &  | 
|
1781  | 
(differ(b::'a,table)) &  | 
|
1782  | 
(differ(c::'a,d)) &  | 
|
1783  | 
(differ(c::'a,table)) &  | 
|
1784  | 
(differ(d::'a,table)) &  | 
|
1785  | 
(holds(on(a::'a,table),s0)) &  | 
|
1786  | 
(holds(on(b::'a,table),s0)) &  | 
|
1787  | 
(holds(on(c::'a,d),s0)) &  | 
|
1788  | 
(holds(on(d::'a,table),s0)) &  | 
|
1789  | 
(holds(clear(a),s0)) &  | 
|
1790  | 
(holds(clear(b),s0)) &  | 
|
1791  | 
(holds(clear(c),s0)) &  | 
|
1792  | 
(holds(EMPTY::'a,s0)) &  | 
|
| 24127 | 1793  | 
(\<forall>State. holds(clear(table),State))"  | 
1794  | 
||
1795  | 
(*190 inferences so far. Searching to depth 10. 0.6 secs*)  | 
|
1796  | 
lemma PLA006_1:  | 
|
1797  | 
"PLA001_0_ax putdown on pickup do holding table differ clear EMPTY and' holds &  | 
|
1798  | 
PLA001_1_ax EMPTY clear s0 on holds table d c b a differ &  | 
|
1799  | 
(\<forall>State. ~holds(on(c::'a,table),State)) --> False"  | 
|
1800  | 
by meson  | 
|
1801  | 
||
1802  | 
(*190 inferences so far. Searching to depth 10. 0.5 secs*)  | 
|
1803  | 
lemma PLA017_1:  | 
|
1804  | 
"PLA001_0_ax putdown on pickup do holding table differ clear EMPTY and' holds &  | 
|
1805  | 
PLA001_1_ax EMPTY clear s0 on holds table d c b a differ &  | 
|
1806  | 
(\<forall>State. ~holds(on(a::'a,c),State)) --> False"  | 
|
1807  | 
by meson  | 
|
1808  | 
||
1809  | 
(*13732 inferences so far. Searching to depth 16. 9.8 secs*)  | 
|
1810  | 
lemma PLA022_1:  | 
|
1811  | 
"PLA001_0_ax putdown on pickup do holding table differ clear EMPTY and' holds &  | 
|
1812  | 
PLA001_1_ax EMPTY clear s0 on holds table d c b a differ &  | 
|
1813  | 
(\<forall>State. ~holds(and'(on(c::'a,d),on(a::'a,c)),State)) --> False"  | 
|
1814  | 
by meson  | 
|
1815  | 
||
1816  | 
(*217 inferences so far. Searching to depth 13. 0.7 secs*)  | 
|
1817  | 
lemma PLA022_2:  | 
|
1818  | 
"PLA001_0_ax putdown on pickup do holding table differ clear EMPTY and' holds &  | 
|
1819  | 
PLA001_1_ax EMPTY clear s0 on holds table d c b a differ &  | 
|
1820  | 
(\<forall>State. ~holds(and'(on(a::'a,c),on(c::'a,d)),State)) --> False"  | 
|
1821  | 
by meson  | 
|
1822  | 
||
1823  | 
(*948 inferences so far. Searching to depth 18. 1.1 secs*)  | 
|
1824  | 
lemma PRV001_1:  | 
|
| 24128 | 1825  | 
"(\<forall>X Y Z. q1(X::'a,Y,Z) & mless_or_equal(X::'a,Y) --> q2(X::'a,Y,Z)) &  | 
1826  | 
(\<forall>X Y Z. q1(X::'a,Y,Z) --> mless_or_equal(X::'a,Y) | q3(X::'a,Y,Z)) &  | 
|
1827  | 
(\<forall>Z X Y. q2(X::'a,Y,Z) --> q4(X::'a,Y,Y)) &  | 
|
1828  | 
(\<forall>Z Y X. q3(X::'a,Y,Z) --> q4(X::'a,Y,X)) &  | 
|
1829  | 
(\<forall>X. mless_or_equal(X::'a,X)) &  | 
|
1830  | 
(\<forall>X Y. mless_or_equal(X::'a,Y) & mless_or_equal(Y::'a,X) --> equal(X::'a,Y)) &  | 
|
1831  | 
(\<forall>Y X Z. mless_or_equal(X::'a,Y) & mless_or_equal(Y::'a,Z) --> mless_or_equal(X::'a,Z)) &  | 
|
1832  | 
(\<forall>Y X. mless_or_equal(X::'a,Y) | mless_or_equal(Y::'a,X)) &  | 
|
1833  | 
(\<forall>X Y. equal(X::'a,Y) --> mless_or_equal(X::'a,Y)) &  | 
|
1834  | 
(\<forall>X Y Z. equal(X::'a,Y) & mless_or_equal(X::'a,Z) --> mless_or_equal(Y::'a,Z)) &  | 
|
1835  | 
(\<forall>X Z Y. equal(X::'a,Y) & mless_or_equal(Z::'a,X) --> mless_or_equal(Z::'a,Y)) &  | 
|
1836  | 
(q1(a::'a,b,c)) &  | 
|
1837  | 
(\<forall>W. ~(q4(a::'a,b,W) & mless_or_equal(a::'a,W) & mless_or_equal(b::'a,W) & mless_or_equal(W::'a,a))) &  | 
|
| 24127 | 1838  | 
(\<forall>W. ~(q4(a::'a,b,W) & mless_or_equal(a::'a,W) & mless_or_equal(b::'a,W) & mless_or_equal(W::'a,b))) --> False"  | 
1839  | 
by meson  | 
|
1840  | 
||
1841  | 
(*PRV is now called SWV (software verification) *)  | 
|
1842  | 
abbreviation "SWV001_1_ax mless_THAN successor predecessor equal \<equiv>  | 
|
| 24128 | 1843  | 
(\<forall>X. equal(predecessor(successor(X)),X)) &  | 
1844  | 
(\<forall>X. equal(successor(predecessor(X)),X)) &  | 
|
1845  | 
(\<forall>X Y. equal(predecessor(X),predecessor(Y)) --> equal(X::'a,Y)) &  | 
|
1846  | 
(\<forall>X Y. equal(successor(X),successor(Y)) --> equal(X::'a,Y)) &  | 
|
1847  | 
(\<forall>X. mless_THAN(predecessor(X),X)) &  | 
|
1848  | 
(\<forall>X. mless_THAN(X::'a,successor(X))) &  | 
|
1849  | 
(\<forall>X Y Z. mless_THAN(X::'a,Y) & mless_THAN(Y::'a,Z) --> mless_THAN(X::'a,Z)) &  | 
|
1850  | 
(\<forall>X Y. mless_THAN(X::'a,Y) | mless_THAN(Y::'a,X) | equal(X::'a,Y)) &  | 
|
1851  | 
(\<forall>X. ~mless_THAN(X::'a,X)) &  | 
|
1852  | 
(\<forall>Y X. ~(mless_THAN(X::'a,Y) & mless_THAN(Y::'a,X))) &  | 
|
1853  | 
(\<forall>Y X Z. equal(X::'a,Y) & mless_THAN(X::'a,Z) --> mless_THAN(Y::'a,Z)) &  | 
|
| 24127 | 1854  | 
(\<forall>Y Z X. equal(X::'a,Y) & mless_THAN(Z::'a,X) --> mless_THAN(Z::'a,Y))"  | 
1855  | 
||
1856  | 
abbreviation "SWV001_0_eq a successor predecessor equal \<equiv>  | 
|
| 24128 | 1857  | 
(\<forall>X Y. equal(X::'a,Y) --> equal(predecessor(X),predecessor(Y))) &  | 
1858  | 
(\<forall>X Y. equal(X::'a,Y) --> equal(successor(X),successor(Y))) &  | 
|
| 24127 | 1859  | 
(\<forall>X Y. equal(X::'a,Y) --> equal(a(X),a(Y)))"  | 
1860  | 
||
1861  | 
(*21 inferences so far. Searching to depth 5. 0.4 secs*)  | 
|
1862  | 
lemma PRV003_1:  | 
|
1863  | 
"EQU001_0_ax equal &  | 
|
1864  | 
SWV001_1_ax mless_THAN successor predecessor equal &  | 
|
| 24128 | 1865  | 
SWV001_0_eq a successor predecessor equal &  | 
1866  | 
(~mless_THAN(n::'a,j)) &  | 
|
1867  | 
(mless_THAN(k::'a,j)) &  | 
|
1868  | 
(~mless_THAN(k::'a,i)) &  | 
|
1869  | 
(mless_THAN(i::'a,n)) &  | 
|
1870  | 
(mless_THAN(a(j),a(k))) &  | 
|
1871  | 
(\<forall>X. mless_THAN(X::'a,j) & mless_THAN(a(X),a(k)) --> mless_THAN(X::'a,i)) &  | 
|
1872  | 
(\<forall>X. mless_THAN(One::'a,i) & mless_THAN(a(X),a(predecessor(i))) --> mless_THAN(X::'a,i) | mless_THAN(n::'a,X)) &  | 
|
1873  | 
(\<forall>X. ~(mless_THAN(One::'a,X) & mless_THAN(X::'a,i) & mless_THAN(a(X),a(predecessor(X))))) &  | 
|
| 24127 | 1874  | 
(mless_THAN(j::'a,i)) --> False"  | 
1875  | 
by meson  | 
|
1876  | 
||
1877  | 
(*584 inferences so far. Searching to depth 7. 1.1 secs*)  | 
|
1878  | 
lemma PRV005_1:  | 
|
1879  | 
"EQU001_0_ax equal &  | 
|
1880  | 
SWV001_1_ax mless_THAN successor predecessor equal &  | 
|
| 24128 | 1881  | 
SWV001_0_eq a successor predecessor equal &  | 
1882  | 
(~mless_THAN(n::'a,k)) &  | 
|
1883  | 
(~mless_THAN(k::'a,l)) &  | 
|
1884  | 
(~mless_THAN(k::'a,i)) &  | 
|
1885  | 
(mless_THAN(l::'a,n)) &  | 
|
1886  | 
(mless_THAN(One::'a,l)) &  | 
|
1887  | 
(mless_THAN(a(k),a(predecessor(l)))) &  | 
|
1888  | 
(\<forall>X. mless_THAN(X::'a,successor(n)) & mless_THAN(a(X),a(k)) --> mless_THAN(X::'a,l)) &  | 
|
1889  | 
(\<forall>X. mless_THAN(One::'a,l) & mless_THAN(a(X),a(predecessor(l))) --> mless_THAN(X::'a,l) | mless_THAN(n::'a,X)) &  | 
|
| 24127 | 1890  | 
(\<forall>X. ~(mless_THAN(One::'a,X) & mless_THAN(X::'a,l) & mless_THAN(a(X),a(predecessor(X))))) --> False"  | 
1891  | 
by meson  | 
|
1892  | 
||
1893  | 
(*2343 inferences so far. Searching to depth 8. 3.5 secs*)  | 
|
1894  | 
lemma PRV006_1:  | 
|
1895  | 
"EQU001_0_ax equal &  | 
|
1896  | 
SWV001_1_ax mless_THAN successor predecessor equal &  | 
|
1897  | 
SWV001_0_eq a successor predecessor equal &  | 
|
| 24128 | 1898  | 
(~mless_THAN(n::'a,m)) &  | 
1899  | 
(mless_THAN(i::'a,m)) &  | 
|
1900  | 
(mless_THAN(i::'a,n)) &  | 
|
1901  | 
(~mless_THAN(i::'a,One)) &  | 
|
1902  | 
(mless_THAN(a(i),a(m))) &  | 
|
1903  | 
(\<forall>X. mless_THAN(X::'a,successor(n)) & mless_THAN(a(X),a(m)) --> mless_THAN(X::'a,i)) &  | 
|
1904  | 
(\<forall>X. mless_THAN(One::'a,i) & mless_THAN(a(X),a(predecessor(i))) --> mless_THAN(X::'a,i) | mless_THAN(n::'a,X)) &  | 
|
| 24127 | 1905  | 
(\<forall>X. ~(mless_THAN(One::'a,X) & mless_THAN(X::'a,i) & mless_THAN(a(X),a(predecessor(X))))) --> False"  | 
1906  | 
by meson  | 
|
1907  | 
||
1908  | 
(*86 inferences so far. Searching to depth 14. 0.1 secs*)  | 
|
1909  | 
lemma PRV009_1:  | 
|
| 24128 | 1910  | 
"(\<forall>Y X. mless_or_equal(X::'a,Y) | mless(Y::'a,X)) &  | 
1911  | 
(mless(j::'a,i)) &  | 
|
1912  | 
(mless_or_equal(m::'a,p)) &  | 
|
1913  | 
(mless_or_equal(p::'a,q)) &  | 
|
1914  | 
(mless_or_equal(q::'a,n)) &  | 
|
1915  | 
(\<forall>X Y. mless_or_equal(m::'a,X) & mless(X::'a,i) & mless(j::'a,Y) & mless_or_equal(Y::'a,n) --> mless_or_equal(a(X),a(Y))) &  | 
|
1916  | 
(\<forall>X Y. mless_or_equal(m::'a,X) & mless_or_equal(X::'a,Y) & mless_or_equal(Y::'a,j) --> mless_or_equal(a(X),a(Y))) &  | 
|
1917  | 
(\<forall>X Y. mless_or_equal(i::'a,X) & mless_or_equal(X::'a,Y) & mless_or_equal(Y::'a,n) --> mless_or_equal(a(X),a(Y))) &  | 
|
| 24127 | 1918  | 
(~mless_or_equal(a(p),a(q))) --> False"  | 
1919  | 
by meson  | 
|
1920  | 
||
1921  | 
(*222 inferences so far. Searching to depth 8. 0.4 secs*)  | 
|
1922  | 
lemma PUZ012_1:  | 
|
| 24128 | 1923  | 
"(\<forall>X. equal_fruits(X::'a,X)) &  | 
1924  | 
(\<forall>X. equal_boxes(X::'a,X)) &  | 
|
1925  | 
(\<forall>X Y. ~(label(X::'a,Y) & contains(X::'a,Y))) &  | 
|
1926  | 
(\<forall>X. contains(boxa::'a,X) | contains(boxb::'a,X) | contains(boxc::'a,X)) &  | 
|
1927  | 
(\<forall>X. contains(X::'a,apples) | contains(X::'a,bananas) | contains(X::'a,oranges)) &  | 
|
1928  | 
(\<forall>X Y Z. contains(X::'a,Y) & contains(X::'a,Z) --> equal_fruits(Y::'a,Z)) &  | 
|
1929  | 
(\<forall>Y X Z. contains(X::'a,Y) & contains(Z::'a,Y) --> equal_boxes(X::'a,Z)) &  | 
|
1930  | 
(~equal_boxes(boxa::'a,boxb)) &  | 
|
1931  | 
(~equal_boxes(boxb::'a,boxc)) &  | 
|
1932  | 
(~equal_boxes(boxa::'a,boxc)) &  | 
|
1933  | 
(~equal_fruits(apples::'a,bananas)) &  | 
|
1934  | 
(~equal_fruits(bananas::'a,oranges)) &  | 
|
1935  | 
(~equal_fruits(apples::'a,oranges)) &  | 
|
1936  | 
(label(boxa::'a,apples)) &  | 
|
1937  | 
(label(boxb::'a,oranges)) &  | 
|
1938  | 
(label(boxc::'a,bananas)) &  | 
|
1939  | 
(contains(boxb::'a,apples)) &  | 
|
| 24127 | 1940  | 
(~(contains(boxa::'a,bananas) & contains(boxc::'a,oranges))) --> False"  | 
1941  | 
by meson  | 
|
1942  | 
||
1943  | 
(*35 inferences so far. Searching to depth 5. 3.2 secs*)  | 
|
1944  | 
lemma PUZ020_1:  | 
|
| 24128 | 1945  | 
"EQU001_0_ax equal &  | 
1946  | 
(\<forall>A B. equal(A::'a,B) --> equal(statement_by(A),statement_by(B))) &  | 
|
1947  | 
(\<forall>X. person(X) --> knight(X) | knave(X)) &  | 
|
1948  | 
(\<forall>X. ~(person(X) & knight(X) & knave(X))) &  | 
|
1949  | 
(\<forall>X Y. says(X::'a,Y) & a_truth(Y) --> a_truth(Y)) &  | 
|
1950  | 
(\<forall>X Y. ~(says(X::'a,Y) & equal(X::'a,Y))) &  | 
|
1951  | 
(\<forall>Y X. says(X::'a,Y) --> equal(Y::'a,statement_by(X))) &  | 
|
1952  | 
(\<forall>X Y. ~(person(X) & equal(X::'a,statement_by(Y)))) &  | 
|
1953  | 
(\<forall>X. person(X) & a_truth(statement_by(X)) --> knight(X)) &  | 
|
1954  | 
(\<forall>X. person(X) --> a_truth(statement_by(X)) | knave(X)) &  | 
|
1955  | 
(\<forall>X Y. equal(X::'a,Y) & knight(X) --> knight(Y)) &  | 
|
1956  | 
(\<forall>X Y. equal(X::'a,Y) & knave(X) --> knave(Y)) &  | 
|
1957  | 
(\<forall>X Y. equal(X::'a,Y) & person(X) --> person(Y)) &  | 
|
1958  | 
(\<forall>X Y Z. equal(X::'a,Y) & says(X::'a,Z) --> says(Y::'a,Z)) &  | 
|
1959  | 
(\<forall>X Z Y. equal(X::'a,Y) & says(Z::'a,X) --> says(Z::'a,Y)) &  | 
|
1960  | 
(\<forall>X Y. equal(X::'a,Y) & a_truth(X) --> a_truth(Y)) &  | 
|
1961  | 
(\<forall>X Y. knight(X) & says(X::'a,Y) --> a_truth(Y)) &  | 
|
1962  | 
(\<forall>X Y. ~(knave(X) & says(X::'a,Y) & a_truth(Y))) &  | 
|
1963  | 
(person(husband)) &  | 
|
1964  | 
(person(wife)) &  | 
|
1965  | 
(~equal(husband::'a,wife)) &  | 
|
1966  | 
(says(husband::'a,statement_by(husband))) &  | 
|
1967  | 
(a_truth(statement_by(husband)) & knight(husband) --> knight(wife)) &  | 
|
1968  | 
(knight(husband) --> a_truth(statement_by(husband))) &  | 
|
1969  | 
(a_truth(statement_by(husband)) | knight(wife)) &  | 
|
1970  | 
(knight(wife) --> a_truth(statement_by(husband))) &  | 
|
| 24127 | 1971  | 
(~knight(husband)) --> False"  | 
1972  | 
by meson  | 
|
1973  | 
||
1974  | 
(*121806 inferences so far. Searching to depth 17. 63.0 secs*)  | 
|
1975  | 
lemma PUZ025_1:  | 
|
| 24128 | 1976  | 
"(\<forall>X. a_truth(truthteller(X)) | a_truth(liar(X))) &  | 
1977  | 
(\<forall>X. ~(a_truth(truthteller(X)) & a_truth(liar(X)))) &  | 
|
1978  | 
(\<forall>Truthteller Statement. a_truth(truthteller(Truthteller)) & a_truth(says(Truthteller::'a,Statement)) --> a_truth(Statement)) &  | 
|
1979  | 
(\<forall>Liar Statement. ~(a_truth(liar(Liar)) & a_truth(says(Liar::'a,Statement)) & a_truth(Statement))) &  | 
|
1980  | 
(\<forall>Statement Truthteller. a_truth(Statement) & a_truth(says(Truthteller::'a,Statement)) --> a_truth(truthteller(Truthteller))) &  | 
|
1981  | 
(\<forall>Statement Liar. a_truth(says(Liar::'a,Statement)) --> a_truth(Statement) | a_truth(liar(Liar))) &  | 
|
1982  | 
(\<forall>Z X Y. people(X::'a,Y,Z) & a_truth(liar(X)) & a_truth(liar(Y)) --> a_truth(equal_type(X::'a,Y))) &  | 
|
1983  | 
(\<forall>Z X Y. people(X::'a,Y,Z) & a_truth(truthteller(X)) & a_truth(truthteller(Y)) --> a_truth(equal_type(X::'a,Y))) &  | 
|
1984  | 
(\<forall>X Y. a_truth(equal_type(X::'a,Y)) & a_truth(truthteller(X)) --> a_truth(truthteller(Y))) &  | 
|
1985  | 
(\<forall>X Y. a_truth(equal_type(X::'a,Y)) & a_truth(liar(X)) --> a_truth(liar(Y))) &  | 
|
1986  | 
(\<forall>X Y. a_truth(truthteller(X)) --> a_truth(equal_type(X::'a,Y)) | a_truth(liar(Y))) &  | 
|
1987  | 
(\<forall>X Y. a_truth(liar(X)) --> a_truth(equal_type(X::'a,Y)) | a_truth(truthteller(Y))) &  | 
|
1988  | 
(\<forall>Y X. a_truth(equal_type(X::'a,Y)) --> a_truth(equal_type(Y::'a,X))) &  | 
|
1989  | 
(\<forall>X Y. ask_1_if_2(X::'a,Y) & a_truth(truthteller(X)) & a_truth(Y) --> answer(yes)) &  | 
|
1990  | 
(\<forall>X Y. ask_1_if_2(X::'a,Y) & a_truth(truthteller(X)) --> a_truth(Y) | answer(no)) &  | 
|
1991  | 
(\<forall>X Y. ask_1_if_2(X::'a,Y) & a_truth(liar(X)) & a_truth(Y) --> answer(no)) &  | 
|
1992  | 
(\<forall>X Y. ask_1_if_2(X::'a,Y) & a_truth(liar(X)) --> a_truth(Y) | answer(yes)) &  | 
|
1993  | 
(people(b::'a,c,a)) &  | 
|
1994  | 
(people(a::'a,b,a)) &  | 
|
1995  | 
(people(a::'a,c,b)) &  | 
|
1996  | 
(people(c::'a,b,a)) &  | 
|
1997  | 
(a_truth(says(a::'a,equal_type(b::'a,c)))) &  | 
|
1998  | 
(ask_1_if_2(c::'a,equal_type(a::'a,b))) &  | 
|
| 24127 | 1999  | 
(\<forall>Answer. ~answer(Answer)) --> False"  | 
2000  | 
oops  | 
|
2001  | 
||
2002  | 
||
2003  | 
(*621 inferences so far. Searching to depth 18. 0.2 secs*)  | 
|
2004  | 
lemma PUZ029_1:  | 
|
| 24128 | 2005  | 
"(\<forall>X. dances_on_tightropes(X) | eats_pennybuns(X) | old(X)) &  | 
2006  | 
(\<forall>X. pig(X) & liable_to_giddiness(X) --> treated_with_respect(X)) &  | 
|
2007  | 
(\<forall>X. wise(X) & balloonist(X) --> has_umbrella(X)) &  | 
|
2008  | 
(\<forall>X. ~(looks_ridiculous(X) & eats_pennybuns(X) & eats_lunch_in_public(X))) &  | 
|
2009  | 
(\<forall>X. balloonist(X) & young(X) --> liable_to_giddiness(X)) &  | 
|
2010  | 
(\<forall>X. fat(X) & looks_ridiculous(X) --> dances_on_tightropes(X) | eats_lunch_in_public(X)) &  | 
|
2011  | 
(\<forall>X. ~(liable_to_giddiness(X) & wise(X) & dances_on_tightropes(X))) &  | 
|
2012  | 
(\<forall>X. pig(X) & has_umbrella(X) --> looks_ridiculous(X)) &  | 
|
2013  | 
(\<forall>X. treated_with_respect(X) --> dances_on_tightropes(X) | fat(X)) &  | 
|
2014  | 
(\<forall>X. young(X) | old(X)) &  | 
|
2015  | 
(\<forall>X. ~(young(X) & old(X))) &  | 
|
2016  | 
(wise(piggy)) &  | 
|
2017  | 
(young(piggy)) &  | 
|
2018  | 
(pig(piggy)) &  | 
|
| 24127 | 2019  | 
(balloonist(piggy)) --> False"  | 
2020  | 
by meson  | 
|
2021  | 
||
2022  | 
abbreviation "RNG001_0_ax equal additive_inverse add multiply product additive_identity sum \<equiv>  | 
|
| 24128 | 2023  | 
(\<forall>X. sum(additive_identity::'a,X,X)) &  | 
2024  | 
(\<forall>X. sum(X::'a,additive_identity,X)) &  | 
|
2025  | 
(\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &  | 
|
2026  | 
(\<forall>X Y. sum(X::'a,Y,add(X::'a,Y))) &  | 
|
2027  | 
(\<forall>X. sum(additive_inverse(X),X,additive_identity)) &  | 
|
2028  | 
(\<forall>X. sum(X::'a,additive_inverse(X),additive_identity)) &  | 
|
2029  | 
(\<forall>Y U Z X V W. sum(X::'a,Y,U) & sum(Y::'a,Z,V) & sum(U::'a,Z,W) --> sum(X::'a,V,W)) &  | 
|
2030  | 
(\<forall>Y X V U Z W. sum(X::'a,Y,U) & sum(Y::'a,Z,V) & sum(X::'a,V,W) --> sum(U::'a,Z,W)) &  | 
|
2031  | 
(\<forall>Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) &  | 
|
2032  | 
(\<forall>Y U Z X V W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(U::'a,Z,W) --> product(X::'a,V,W)) &  | 
|
2033  | 
(\<forall>Y X V U Z W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(X::'a,V,W) --> product(U::'a,Z,W)) &  | 
|
2034  | 
(\<forall>Y Z X V3 V1 V2 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & product(X::'a,V3,V4) --> sum(V1::'a,V2,V4)) &  | 
|
2035  | 
(\<forall>Y Z V1 V2 X V3 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & sum(V1::'a,V2,V4) --> product(X::'a,V3,V4)) &  | 
|
2036  | 
(\<forall>Y Z V3 X V1 V2 V4. product(Y::'a,X,V1) & product(Z::'a,X,V2) & sum(Y::'a,Z,V3) & product(V3::'a,X,V4) --> sum(V1::'a,V2,V4)) &  | 
|
2037  | 
(\<forall>Y Z V1 V2 V3 X V4. product(Y::'a,X,V1) & product(Z::'a,X,V2) & sum(Y::'a,Z,V3) & sum(V1::'a,V2,V4) --> product(V3::'a,X,V4)) &  | 
|
2038  | 
(\<forall>X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) &  | 
|
| 24127 | 2039  | 
(\<forall>X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V))"  | 
2040  | 
||
2041  | 
abbreviation "RNG001_0_eq product multiply sum add additive_inverse equal \<equiv>  | 
|
| 24128 | 2042  | 
(\<forall>X Y. equal(X::'a,Y) --> equal(additive_inverse(X),additive_inverse(Y))) &  | 
2043  | 
(\<forall>X Y W. equal(X::'a,Y) --> equal(add(X::'a,W),add(Y::'a,W))) &  | 
|
2044  | 
(\<forall>X W Y. equal(X::'a,Y) --> equal(add(W::'a,X),add(W::'a,Y))) &  | 
|
2045  | 
(\<forall>X Y W Z. equal(X::'a,Y) & sum(X::'a,W,Z) --> sum(Y::'a,W,Z)) &  | 
|
2046  | 
(\<forall>X W Y Z. equal(X::'a,Y) & sum(W::'a,X,Z) --> sum(W::'a,Y,Z)) &  | 
|
2047  | 
(\<forall>X W Z Y. equal(X::'a,Y) & sum(W::'a,Z,X) --> sum(W::'a,Z,Y)) &  | 
|
2048  | 
(\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &  | 
|
2049  | 
(\<forall>X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) &  | 
|
2050  | 
(\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &  | 
|
2051  | 
(\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &  | 
|
| 24127 | 2052  | 
(\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y))"  | 
2053  | 
||
2054  | 
(*93620 inferences so far. Searching to depth 24. 65.9 secs*)  | 
|
2055  | 
lemma RNG001_3:  | 
|
| 24128 | 2056  | 
"(\<forall>X. sum(additive_identity::'a,X,X)) &  | 
2057  | 
(\<forall>X. sum(additive_inverse(X),X,additive_identity)) &  | 
|
2058  | 
(\<forall>Y U Z X V W. sum(X::'a,Y,U) & sum(Y::'a,Z,V) & sum(U::'a,Z,W) --> sum(X::'a,V,W)) &  | 
|
2059  | 
(\<forall>Y X V U Z W. sum(X::'a,Y,U) & sum(Y::'a,Z,V) & sum(X::'a,V,W) --> sum(U::'a,Z,W)) &  | 
|
2060  | 
(\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &  | 
|
2061  | 
(\<forall>Y Z X V3 V1 V2 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & product(X::'a,V3,V4) --> sum(V1::'a,V2,V4)) &  | 
|
2062  | 
(\<forall>Y Z V1 V2 X V3 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & sum(V1::'a,V2,V4) --> product(X::'a,V3,V4)) &  | 
|
| 24127 | 2063  | 
(~product(a::'a,additive_identity,additive_identity)) --> False"  | 
2064  | 
oops  | 
|
2065  | 
||
2066  | 
abbreviation "RNG_other_ax multiply add equal product additive_identity additive_inverse sum \<equiv>  | 
|
| 24128 | 2067  | 
(\<forall>X. sum(X::'a,additive_inverse(X),additive_identity)) &  | 
2068  | 
(\<forall>Y U Z X V W. sum(X::'a,Y,U) & sum(Y::'a,Z,V) & sum(U::'a,Z,W) --> sum(X::'a,V,W)) &  | 
|
2069  | 
(\<forall>Y X V U Z W. sum(X::'a,Y,U) & sum(Y::'a,Z,V) & sum(X::'a,V,W) --> sum(U::'a,Z,W)) &  | 
|
2070  | 
(\<forall>Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) &  | 
|
2071  | 
(\<forall>Y U Z X V W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(U::'a,Z,W) --> product(X::'a,V,W)) &  | 
|
2072  | 
(\<forall>Y X V U Z W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(X::'a,V,W) --> product(U::'a,Z,W)) &  | 
|
2073  | 
(\<forall>Y Z X V3 V1 V2 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & product(X::'a,V3,V4) --> sum(V1::'a,V2,V4)) &  | 
|
2074  | 
(\<forall>Y Z V1 V2 X V3 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & sum(V1::'a,V2,V4) --> product(X::'a,V3,V4)) &  | 
|
2075  | 
(\<forall>Y Z V3 X V1 V2 V4. product(Y::'a,X,V1) & product(Z::'a,X,V2) & sum(Y::'a,Z,V3) & product(V3::'a,X,V4) --> sum(V1::'a,V2,V4)) &  | 
|
2076  | 
(\<forall>Y Z V1 V2 V3 X V4. product(Y::'a,X,V1) & product(Z::'a,X,V2) & sum(Y::'a,Z,V3) & sum(V1::'a,V2,V4) --> product(V3::'a,X,V4)) &  | 
|
2077  | 
(\<forall>X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) &  | 
|
2078  | 
(\<forall>X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V)) &  | 
|
2079  | 
(\<forall>X Y. equal(X::'a,Y) --> equal(additive_inverse(X),additive_inverse(Y))) &  | 
|
2080  | 
(\<forall>X Y W. equal(X::'a,Y) --> equal(add(X::'a,W),add(Y::'a,W))) &  | 
|
2081  | 
(\<forall>X Y W Z. equal(X::'a,Y) & sum(X::'a,W,Z) --> sum(Y::'a,W,Z)) &  | 
|
2082  | 
(\<forall>X W Y Z. equal(X::'a,Y) & sum(W::'a,X,Z) --> sum(W::'a,Y,Z)) &  | 
|
2083  | 
(\<forall>X W Z Y. equal(X::'a,Y) & sum(W::'a,Z,X) --> sum(W::'a,Z,Y)) &  | 
|
2084  | 
(\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &  | 
|
2085  | 
(\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &  | 
|
2086  | 
(\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &  | 
|
| 24127 | 2087  | 
(\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y))"  | 
2088  | 
||
2089  | 
||
2090  | 
(****************SLOW  | 
|
2091  | 
76385914 inferences so far. Searching to depth 18  | 
|
2092  | 
No proof after 5 1/2 hours! (griffon)  | 
|
2093  | 
****************)  | 
|
| 24128 | 2094  | 
lemma RNG001_5:  | 
2095  | 
"EQU001_0_ax equal &  | 
|
2096  | 
(\<forall>X. sum(additive_identity::'a,X,X)) &  | 
|
2097  | 
(\<forall>X. sum(X::'a,additive_identity,X)) &  | 
|
2098  | 
(\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &  | 
|
2099  | 
(\<forall>X Y. sum(X::'a,Y,add(X::'a,Y))) &  | 
|
2100  | 
(\<forall>X. sum(additive_inverse(X),X,additive_identity)) &  | 
|
2101  | 
RNG_other_ax multiply add equal product additive_identity additive_inverse sum &  | 
|
2102  | 
(~product(a::'a,additive_identity,additive_identity)) --> False"  | 
|
2103  | 
oops  | 
|
| 24127 | 2104  | 
|
2105  | 
(*0 inferences so far. Searching to depth 0. 0.5 secs*)  | 
|
2106  | 
lemma RNG011_5:  | 
|
| 24128 | 2107  | 
"EQU001_0_ax equal &  | 
2108  | 
(\<forall>A B C. equal(A::'a,B) --> equal(add(A::'a,C),add(B::'a,C))) &  | 
|
2109  | 
(\<forall>D F' E. equal(D::'a,E) --> equal(add(F'::'a,D),add(F'::'a,E))) &  | 
|
2110  | 
(\<forall>G H. equal(G::'a,H) --> equal(additive_inverse(G),additive_inverse(H))) &  | 
|
2111  | 
(\<forall>I' J K'. equal(I'::'a,J) --> equal(multiply(I'::'a,K'),multiply(J::'a,K'))) &  | 
|
2112  | 
(\<forall>L N M. equal(L::'a,M) --> equal(multiply(N::'a,L),multiply(N::'a,M))) &  | 
|
2113  | 
(\<forall>A B C D. equal(A::'a,B) --> equal(associator(A::'a,C,D),associator(B::'a,C,D))) &  | 
|
2114  | 
(\<forall>E G F' H. equal(E::'a,F') --> equal(associator(G::'a,E,H),associator(G::'a,F',H))) &  | 
|
2115  | 
(\<forall>I' K' L J. equal(I'::'a,J) --> equal(associator(K'::'a,L,I'),associator(K'::'a,L,J))) &  | 
|
2116  | 
(\<forall>M N O'. equal(M::'a,N) --> equal(commutator(M::'a,O'),commutator(N::'a,O'))) &  | 
|
2117  | 
(\<forall>P R Q. equal(P::'a,Q) --> equal(commutator(R::'a,P),commutator(R::'a,Q))) &  | 
|
2118  | 
(\<forall>Y X. equal(add(X::'a,Y),add(Y::'a,X))) &  | 
|
2119  | 
(\<forall>X Y Z. equal(add(add(X::'a,Y),Z),add(X::'a,add(Y::'a,Z)))) &  | 
|
2120  | 
(\<forall>X. equal(add(X::'a,additive_identity),X)) &  | 
|
2121  | 
(\<forall>X. equal(add(additive_identity::'a,X),X)) &  | 
|
2122  | 
(\<forall>X. equal(add(X::'a,additive_inverse(X)),additive_identity)) &  | 
|
2123  | 
(\<forall>X. equal(add(additive_inverse(X),X),additive_identity)) &  | 
|
2124  | 
(equal(additive_inverse(additive_identity),additive_identity)) &  | 
|
2125  | 
(\<forall>X Y. equal(add(X::'a,add(additive_inverse(X),Y)),Y)) &  | 
|
2126  | 
(\<forall>X Y. equal(additive_inverse(add(X::'a,Y)),add(additive_inverse(X),additive_inverse(Y)))) &  | 
|
2127  | 
(\<forall>X. equal(additive_inverse(additive_inverse(X)),X)) &  | 
|
2128  | 
(\<forall>X. equal(multiply(X::'a,additive_identity),additive_identity)) &  | 
|
2129  | 
(\<forall>X. equal(multiply(additive_identity::'a,X),additive_identity)) &  | 
|
2130  | 
(\<forall>X Y. equal(multiply(additive_inverse(X),additive_inverse(Y)),multiply(X::'a,Y))) &  | 
|
2131  | 
(\<forall>X Y. equal(multiply(X::'a,additive_inverse(Y)),additive_inverse(multiply(X::'a,Y)))) &  | 
|
2132  | 
(\<forall>X Y. equal(multiply(additive_inverse(X),Y),additive_inverse(multiply(X::'a,Y)))) &  | 
|
2133  | 
(\<forall>Y X Z. equal(multiply(X::'a,add(Y::'a,Z)),add(multiply(X::'a,Y),multiply(X::'a,Z)))) &  | 
|
2134  | 
(\<forall>X Y Z. equal(multiply(add(X::'a,Y),Z),add(multiply(X::'a,Z),multiply(Y::'a,Z)))) &  | 
|
2135  | 
(\<forall>X Y. equal(multiply(multiply(X::'a,Y),Y),multiply(X::'a,multiply(Y::'a,Y)))) &  | 
|
2136  | 
(\<forall>X Y Z. equal(associator(X::'a,Y,Z),add(multiply(multiply(X::'a,Y),Z),additive_inverse(multiply(X::'a,multiply(Y::'a,Z)))))) &  | 
|
2137  | 
(\<forall>X Y. equal(commutator(X::'a,Y),add(multiply(Y::'a,X),additive_inverse(multiply(X::'a,Y))))) &  | 
|
2138  | 
(\<forall>X Y. equal(multiply(multiply(associator(X::'a,X,Y),X),associator(X::'a,X,Y)),additive_identity)) &  | 
|
| 24127 | 2139  | 
(~equal(multiply(multiply(associator(a::'a,a,b),a),associator(a::'a,a,b)),additive_identity)) --> False"  | 
2140  | 
by meson  | 
|
2141  | 
||
2142  | 
(*202 inferences so far. Searching to depth 8. 0.6 secs*)  | 
|
2143  | 
lemma RNG023_6:  | 
|
| 24128 | 2144  | 
"EQU001_0_ax equal &  | 
2145  | 
(\<forall>Y X. equal(add(X::'a,Y),add(Y::'a,X))) &  | 
|
2146  | 
(\<forall>X Y Z. equal(add(X::'a,add(Y::'a,Z)),add(add(X::'a,Y),Z))) &  | 
|
2147  | 
(\<forall>X. equal(add(additive_identity::'a,X),X)) &  | 
|
2148  | 
(\<forall>X. equal(add(X::'a,additive_identity),X)) &  | 
|
2149  | 
(\<forall>X. equal(multiply(additive_identity::'a,X),additive_identity)) &  | 
|
2150  | 
(\<forall>X. equal(multiply(X::'a,additive_identity),additive_identity)) &  | 
|
2151  | 
(\<forall>X. equal(add(additive_inverse(X),X),additive_identity)) &  | 
|
2152  | 
(\<forall>X. equal(add(X::'a,additive_inverse(X)),additive_identity)) &  | 
|
2153  | 
(\<forall>Y X Z. equal(multiply(X::'a,add(Y::'a,Z)),add(multiply(X::'a,Y),multiply(X::'a,Z)))) &  | 
|
2154  | 
(\<forall>X Y Z. equal(multiply(add(X::'a,Y),Z),add(multiply(X::'a,Z),multiply(Y::'a,Z)))) &  | 
|
2155  | 
(\<forall>X. equal(additive_inverse(additive_inverse(X)),X)) &  | 
|
2156  | 
(\<forall>X Y. equal(multiply(multiply(X::'a,Y),Y),multiply(X::'a,multiply(Y::'a,Y)))) &  | 
|
2157  | 
(\<forall>X Y. equal(multiply(multiply(X::'a,X),Y),multiply(X::'a,multiply(X::'a,Y)))) &  | 
|
2158  | 
(\<forall>X Y Z. equal(associator(X::'a,Y,Z),add(multiply(multiply(X::'a,Y),Z),additive_inverse(multiply(X::'a,multiply(Y::'a,Z)))))) &  | 
|
2159  | 
(\<forall>X Y. equal(commutator(X::'a,Y),add(multiply(Y::'a,X),additive_inverse(multiply(X::'a,Y))))) &  | 
|
2160  | 
(\<forall>D E F'. equal(D::'a,E) --> equal(add(D::'a,F'),add(E::'a,F'))) &  | 
|
2161  | 
(\<forall>G I' H. equal(G::'a,H) --> equal(add(I'::'a,G),add(I'::'a,H))) &  | 
|
2162  | 
(\<forall>J K'. equal(J::'a,K') --> equal(additive_inverse(J),additive_inverse(K'))) &  | 
|
2163  | 
(\<forall>L M N O'. equal(L::'a,M) --> equal(associator(L::'a,N,O'),associator(M::'a,N,O'))) &  | 
|
2164  | 
(\<forall>P R Q S'. equal(P::'a,Q) --> equal(associator(R::'a,P,S'),associator(R::'a,Q,S'))) &  | 
|
2165  | 
(\<forall>T' V W U. equal(T'::'a,U) --> equal(associator(V::'a,W,T'),associator(V::'a,W,U))) &  | 
|
2166  | 
(\<forall>X Y Z. equal(X::'a,Y) --> equal(commutator(X::'a,Z),commutator(Y::'a,Z))) &  | 
|
2167  | 
(\<forall>A1 C1 B1. equal(A1::'a,B1) --> equal(commutator(C1::'a,A1),commutator(C1::'a,B1))) &  | 
|
2168  | 
(\<forall>D1 E1 F1. equal(D1::'a,E1) --> equal(multiply(D1::'a,F1),multiply(E1::'a,F1))) &  | 
|
2169  | 
(\<forall>G1 I1 H1. equal(G1::'a,H1) --> equal(multiply(I1::'a,G1),multiply(I1::'a,H1))) &  | 
|
| 24127 | 2170  | 
(~equal(associator(x::'a,x,y),additive_identity)) --> False"  | 
2171  | 
by meson  | 
|
2172  | 
||
2173  | 
(*0 inferences so far. Searching to depth 0. 0.6 secs*)  | 
|
2174  | 
lemma RNG028_2:  | 
|
| 24128 | 2175  | 
"EQU001_0_ax equal &  | 
2176  | 
(\<forall>X. equal(add(additive_identity::'a,X),X)) &  | 
|
2177  | 
(\<forall>X. equal(multiply(additive_identity::'a,X),additive_identity)) &  | 
|
2178  | 
(\<forall>X. equal(multiply(X::'a,additive_identity),additive_identity)) &  | 
|
2179  | 
(\<forall>X. equal(add(additive_inverse(X),X),additive_identity)) &  | 
|
2180  | 
(\<forall>X Y. equal(additive_inverse(add(X::'a,Y)),add(additive_inverse(X),additive_inverse(Y)))) &  | 
|
2181  | 
(\<forall>X. equal(additive_inverse(additive_inverse(X)),X)) &  | 
|
2182  | 
(\<forall>Y X Z. equal(multiply(X::'a,add(Y::'a,Z)),add(multiply(X::'a,Y),multiply(X::'a,Z)))) &  | 
|
2183  | 
(\<forall>X Y Z. equal(multiply(add(X::'a,Y),Z),add(multiply(X::'a,Z),multiply(Y::'a,Z)))) &  | 
|
2184  | 
(\<forall>X Y. equal(multiply(multiply(X::'a,Y),Y),multiply(X::'a,multiply(Y::'a,Y)))) &  | 
|
2185  | 
(\<forall>X Y. equal(multiply(multiply(X::'a,X),Y),multiply(X::'a,multiply(X::'a,Y)))) &  | 
|
2186  | 
(\<forall>X Y. equal(multiply(additive_inverse(X),Y),additive_inverse(multiply(X::'a,Y)))) &  | 
|
2187  | 
(\<forall>X Y. equal(multiply(X::'a,additive_inverse(Y)),additive_inverse(multiply(X::'a,Y)))) &  | 
|
2188  | 
(equal(additive_inverse(additive_identity),additive_identity)) &  | 
|
2189  | 
(\<forall>Y X. equal(add(X::'a,Y),add(Y::'a,X))) &  | 
|
2190  | 
(\<forall>X Y Z. equal(add(X::'a,add(Y::'a,Z)),add(add(X::'a,Y),Z))) &  | 
|
2191  | 
(\<forall>Z X Y. equal(add(X::'a,Z),add(Y::'a,Z)) --> equal(X::'a,Y)) &  | 
|
2192  | 
(\<forall>Z X Y. equal(add(Z::'a,X),add(Z::'a,Y)) --> equal(X::'a,Y)) &  | 
|
2193  | 
(\<forall>D E F'. equal(D::'a,E) --> equal(add(D::'a,F'),add(E::'a,F'))) &  | 
|
2194  | 
(\<forall>G I' H. equal(G::'a,H) --> equal(add(I'::'a,G),add(I'::'a,H))) &  | 
|
2195  | 
(\<forall>J K'. equal(J::'a,K') --> equal(additive_inverse(J),additive_inverse(K'))) &  | 
|
2196  | 
(\<forall>D1 E1 F1. equal(D1::'a,E1) --> equal(multiply(D1::'a,F1),multiply(E1::'a,F1))) &  | 
|
2197  | 
(\<forall>G1 I1 H1. equal(G1::'a,H1) --> equal(multiply(I1::'a,G1),multiply(I1::'a,H1))) &  | 
|
2198  | 
(\<forall>X Y Z. equal(associator(X::'a,Y,Z),add(multiply(multiply(X::'a,Y),Z),additive_inverse(multiply(X::'a,multiply(Y::'a,Z)))))) &  | 
|
2199  | 
(\<forall>L M N O'. equal(L::'a,M) --> equal(associator(L::'a,N,O'),associator(M::'a,N,O'))) &  | 
|
2200  | 
(\<forall>P R Q S'. equal(P::'a,Q) --> equal(associator(R::'a,P,S'),associator(R::'a,Q,S'))) &  | 
|
2201  | 
(\<forall>T' V W U. equal(T'::'a,U) --> equal(associator(V::'a,W,T'),associator(V::'a,W,U))) &  | 
|
2202  | 
(\<forall>X Y. ~equal(multiply(multiply(Y::'a,X),Y),multiply(Y::'a,multiply(X::'a,Y)))) &  | 
|
2203  | 
(\<forall>X Y Z. ~equal(associator(Y::'a,X,Z),additive_inverse(associator(X::'a,Y,Z)))) &  | 
|
2204  | 
(\<forall>X Y Z. ~equal(associator(Z::'a,Y,X),additive_inverse(associator(X::'a,Y,Z)))) &  | 
|
| 24127 | 2205  | 
(~equal(multiply(multiply(cx::'a,multiply(cy::'a,cx)),cz),multiply(cx::'a,multiply(cy::'a,multiply(cx::'a,cz))))) --> False"  | 
2206  | 
by meson  | 
|
2207  | 
||
2208  | 
(*209 inferences so far. Searching to depth 9. 1.2 secs*)  | 
|
2209  | 
lemma RNG038_2:  | 
|
| 24128 | 2210  | 
"(\<forall>X. sum(X::'a,additive_identity,X)) &  | 
2211  | 
(\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &  | 
|
| 24127 | 2212  | 
(\<forall>X Y. sum(X::'a,Y,add(X::'a,Y))) &  | 
| 24128 | 2213  | 
RNG_other_ax multiply add equal product additive_identity additive_inverse sum &  | 
2214  | 
(\<forall>X. product(additive_identity::'a,X,additive_identity)) &  | 
|
2215  | 
(\<forall>X. product(X::'a,additive_identity,additive_identity)) &  | 
|
2216  | 
(\<forall>X Y. equal(X::'a,additive_identity) --> product(X::'a,h(X::'a,Y),Y)) &  | 
|
2217  | 
(product(a::'a,b,additive_identity)) &  | 
|
2218  | 
(~equal(a::'a,additive_identity)) &  | 
|
| 24127 | 2219  | 
(~equal(b::'a,additive_identity)) --> False"  | 
2220  | 
by meson  | 
|
2221  | 
||
2222  | 
(*2660 inferences so far. Searching to depth 10. 7.0 secs*)  | 
|
2223  | 
lemma RNG040_2:  | 
|
2224  | 
"EQU001_0_ax equal &  | 
|
| 24128 | 2225  | 
RNG001_0_eq product multiply sum add additive_inverse equal &  | 
2226  | 
(\<forall>X. sum(additive_identity::'a,X,X)) &  | 
|
2227  | 
(\<forall>X. sum(X::'a,additive_identity,X)) &  | 
|
2228  | 
(\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &  | 
|
2229  | 
(\<forall>X Y. sum(X::'a,Y,add(X::'a,Y))) &  | 
|
2230  | 
(\<forall>X. sum(additive_inverse(X),X,additive_identity)) &  | 
|
2231  | 
(\<forall>X. sum(X::'a,additive_inverse(X),additive_identity)) &  | 
|
2232  | 
(\<forall>Y U Z X V W. sum(X::'a,Y,U) & sum(Y::'a,Z,V) & sum(U::'a,Z,W) --> sum(X::'a,V,W)) &  | 
|
2233  | 
(\<forall>Y X V U Z W. sum(X::'a,Y,U) & sum(Y::'a,Z,V) & sum(X::'a,V,W) --> sum(U::'a,Z,W)) &  | 
|
2234  | 
(\<forall>Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) &  | 
|
2235  | 
(\<forall>Y U Z X V W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(U::'a,Z,W) --> product(X::'a,V,W)) &  | 
|
2236  | 
(\<forall>Y X V U Z W. product(X::'a,Y,U) & product(Y::'a,Z,V) & product(X::'a,V,W) --> product(U::'a,Z,W)) &  | 
|
2237  | 
(\<forall>Y Z X V3 V1 V2 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & product(X::'a,V3,V4) --> sum(V1::'a,V2,V4)) &  | 
|
2238  | 
(\<forall>Y Z V1 V2 X V3 V4. product(X::'a,Y,V1) & product(X::'a,Z,V2) & sum(Y::'a,Z,V3) & sum(V1::'a,V2,V4) --> product(X::'a,V3,V4)) &  | 
|
2239  | 
(\<forall>X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) &  | 
|
2240  | 
(\<forall>X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V)) &  | 
|
2241  | 
(\<forall>A. product(A::'a,multiplicative_identity,A)) &  | 
|
2242  | 
(\<forall>A. product(multiplicative_identity::'a,A,A)) &  | 
|
2243  | 
(\<forall>A. product(A::'a,h(A),multiplicative_identity) | equal(A::'a,additive_identity)) &  | 
|
2244  | 
(\<forall>A. product(h(A),A,multiplicative_identity) | equal(A::'a,additive_identity)) &  | 
|
2245  | 
(\<forall>B A C. product(A::'a,B,C) --> product(B::'a,A,C)) &  | 
|
2246  | 
(\<forall>A B. equal(A::'a,B) --> equal(h(A),h(B))) &  | 
|
2247  | 
(sum(b::'a,c,d)) &  | 
|
2248  | 
(product(d::'a,a,additive_identity)) &  | 
|
2249  | 
(product(b::'a,a,l)) &  | 
|
2250  | 
(product(c::'a,a,n)) &  | 
|
| 24127 | 2251  | 
(~sum(l::'a,n,additive_identity)) --> False"  | 
2252  | 
by meson  | 
|
2253  | 
||
2254  | 
(*8991 inferences so far. Searching to depth 9. 22.2 secs*)  | 
|
2255  | 
lemma RNG041_1:  | 
|
2256  | 
"EQU001_0_ax equal &  | 
|
2257  | 
RNG001_0_ax equal additive_inverse add multiply product additive_identity sum &  | 
|
| 24128 | 2258  | 
RNG001_0_eq product multiply sum add additive_inverse equal &  | 
2259  | 
(\<forall>A B. equal(A::'a,B) --> equal(h(A),h(B))) &  | 
|
2260  | 
(\<forall>A. product(additive_identity::'a,A,additive_identity)) &  | 
|
2261  | 
(\<forall>A. product(A::'a,additive_identity,additive_identity)) &  | 
|
2262  | 
(\<forall>A. product(A::'a,multiplicative_identity,A)) &  | 
|
2263  | 
(\<forall>A. product(multiplicative_identity::'a,A,A)) &  | 
|
2264  | 
(\<forall>A. product(A::'a,h(A),multiplicative_identity) | equal(A::'a,additive_identity)) &  | 
|
2265  | 
(\<forall>A. product(h(A),A,multiplicative_identity) | equal(A::'a,additive_identity)) &  | 
|
2266  | 
(product(a::'a,b,additive_identity)) &  | 
|
2267  | 
(~equal(a::'a,additive_identity)) &  | 
|
| 24127 | 2268  | 
(~equal(b::'a,additive_identity)) --> False"  | 
2269  | 
oops  | 
|
2270  | 
||
2271  | 
(*101319 inferences so far. Searching to depth 14. 76.0 secs*)  | 
|
2272  | 
lemma ROB010_1:  | 
|
| 24128 | 2273  | 
"EQU001_0_ax equal &  | 
2274  | 
(\<forall>Y X. equal(add(X::'a,Y),add(Y::'a,X))) &  | 
|
2275  | 
(\<forall>X Y Z. equal(add(add(X::'a,Y),Z),add(X::'a,add(Y::'a,Z)))) &  | 
|
2276  | 
(\<forall>Y X. equal(negate(add(negate(add(X::'a,Y)),negate(add(X::'a,negate(Y))))),X)) &  | 
|
2277  | 
(\<forall>A B C. equal(A::'a,B) --> equal(add(A::'a,C),add(B::'a,C))) &  | 
|
2278  | 
(\<forall>D F' E. equal(D::'a,E) --> equal(add(F'::'a,D),add(F'::'a,E))) &  | 
|
2279  | 
(\<forall>G H. equal(G::'a,H) --> equal(negate(G),negate(H))) &  | 
|
2280  | 
(equal(negate(add(a::'a,negate(b))),c)) &  | 
|
| 24127 | 2281  | 
(~equal(negate(add(c::'a,negate(add(b::'a,a)))),a)) --> False"  | 
2282  | 
oops  | 
|
2283  | 
||
2284  | 
||
2285  | 
(*6933 inferences so far. Searching to depth 12. 5.1 secs*)  | 
|
2286  | 
lemma ROB013_1:  | 
|
2287  | 
"EQU001_0_ax equal &  | 
|
| 24128 | 2288  | 
(\<forall>Y X. equal(add(X::'a,Y),add(Y::'a,X))) &  | 
2289  | 
(\<forall>X Y Z. equal(add(add(X::'a,Y),Z),add(X::'a,add(Y::'a,Z)))) &  | 
|
2290  | 
(\<forall>Y X. equal(negate(add(negate(add(X::'a,Y)),negate(add(X::'a,negate(Y))))),X)) &  | 
|
2291  | 
(\<forall>A B C. equal(A::'a,B) --> equal(add(A::'a,C),add(B::'a,C))) &  | 
|
2292  | 
(\<forall>D F' E. equal(D::'a,E) --> equal(add(F'::'a,D),add(F'::'a,E))) &  | 
|
2293  | 
(\<forall>G H. equal(G::'a,H) --> equal(negate(G),negate(H))) &  | 
|
2294  | 
(equal(negate(add(a::'a,b)),c)) &  | 
|
| 24127 | 2295  | 
(~equal(negate(add(c::'a,negate(add(negate(b),a)))),a)) --> False"  | 
2296  | 
by meson  | 
|
2297  | 
||
2298  | 
(*6614 inferences so far. Searching to depth 11. 20.4 secs*)  | 
|
2299  | 
lemma ROB016_1:  | 
|
2300  | 
"EQU001_0_ax equal &  | 
|
| 24128 | 2301  | 
(\<forall>Y X. equal(add(X::'a,Y),add(Y::'a,X))) &  | 
2302  | 
(\<forall>X Y Z. equal(add(add(X::'a,Y),Z),add(X::'a,add(Y::'a,Z)))) &  | 
|
2303  | 
(\<forall>Y X. equal(negate(add(negate(add(X::'a,Y)),negate(add(X::'a,negate(Y))))),X)) &  | 
|
2304  | 
(\<forall>A B C. equal(A::'a,B) --> equal(add(A::'a,C),add(B::'a,C))) &  | 
|
2305  | 
(\<forall>D F' E. equal(D::'a,E) --> equal(add(F'::'a,D),add(F'::'a,E))) &  | 
|
2306  | 
(\<forall>G H. equal(G::'a,H) --> equal(negate(G),negate(H))) &  | 
|
2307  | 
(\<forall>J K' L. equal(J::'a,K') --> equal(multiply(J::'a,L),multiply(K'::'a,L))) &  | 
|
2308  | 
(\<forall>M O' N. equal(M::'a,N) --> equal(multiply(O'::'a,M),multiply(O'::'a,N))) &  | 
|
2309  | 
(\<forall>P Q. equal(P::'a,Q) --> equal(successor(P),successor(Q))) &  | 
|
2310  | 
(\<forall>R S'. equal(R::'a,S') & positive_integer(R) --> positive_integer(S')) &  | 
|
2311  | 
(\<forall>X. equal(multiply(One::'a,X),X)) &  | 
|
2312  | 
(\<forall>V X. positive_integer(X) --> equal(multiply(successor(V),X),add(X::'a,multiply(V::'a,X)))) &  | 
|
2313  | 
(positive_integer(One)) &  | 
|
2314  | 
(\<forall>X. positive_integer(X) --> positive_integer(successor(X))) &  | 
|
2315  | 
(equal(negate(add(d::'a,e)),negate(e))) &  | 
|
2316  | 
(positive_integer(k)) &  | 
|
2317  | 
(\<forall>Vk X Y. equal(negate(add(negate(Y),negate(add(X::'a,negate(Y))))),X) & positive_integer(Vk) --> equal(negate(add(Y::'a,multiply(Vk::'a,add(X::'a,negate(add(X::'a,negate(Y))))))),negate(Y))) &  | 
|
| 24127 | 2318  | 
(~equal(negate(add(e::'a,multiply(k::'a,add(d::'a,negate(add(d::'a,negate(e))))))),negate(e))) --> False"  | 
2319  | 
oops  | 
|
2320  | 
||
2321  | 
(*14077 inferences so far. Searching to depth 11. 32.8 secs*)  | 
|
2322  | 
lemma ROB021_1:  | 
|
2323  | 
"EQU001_0_ax equal &  | 
|
| 24128 | 2324  | 
(\<forall>Y X. equal(add(X::'a,Y),add(Y::'a,X))) &  | 
2325  | 
(\<forall>X Y Z. equal(add(add(X::'a,Y),Z),add(X::'a,add(Y::'a,Z)))) &  | 
|
2326  | 
(\<forall>Y X. equal(negate(add(negate(add(X::'a,Y)),negate(add(X::'a,negate(Y))))),X)) &  | 
|
2327  | 
(\<forall>A B C. equal(A::'a,B) --> equal(add(A::'a,C),add(B::'a,C))) &  | 
|
2328  | 
(\<forall>D F' E. equal(D::'a,E) --> equal(add(F'::'a,D),add(F'::'a,E))) &  | 
|
2329  | 
(\<forall>G H. equal(G::'a,H) --> equal(negate(G),negate(H))) &  | 
|
2330  | 
(\<forall>X Y. equal(negate(X),negate(Y)) --> equal(X::'a,Y)) &  | 
|
| 24127 | 2331  | 
(~equal(add(negate(add(a::'a,negate(b))),negate(add(negate(a),negate(b)))),b)) --> False"  | 
2332  | 
oops  | 
|
2333  | 
||
2334  | 
(*35532 inferences so far. Searching to depth 19. 54.3 secs*)  | 
|
2335  | 
lemma SET005_1:  | 
|
| 24128 | 2336  | 
"(\<forall>Subset Element Superset. member(Element::'a,Subset) & subset(Subset::'a,Superset) --> member(Element::'a,Superset)) &  | 
2337  | 
(\<forall>Superset Subset. subset(Subset::'a,Superset) | member(member_of_1_not_of_2(Subset::'a,Superset),Subset)) &  | 
|
2338  | 
(\<forall>Subset Superset. member(member_of_1_not_of_2(Subset::'a,Superset),Superset) --> subset(Subset::'a,Superset)) &  | 
|
2339  | 
(\<forall>Subset Superset. equal_sets(Subset::'a,Superset) --> subset(Subset::'a,Superset)) &  | 
|
2340  | 
(\<forall>Subset Superset. equal_sets(Superset::'a,Subset) --> subset(Subset::'a,Superset)) &  | 
|
2341  | 
(\<forall>Set2 Set1. subset(Set1::'a,Set2) & subset(Set2::'a,Set1) --> equal_sets(Set2::'a,Set1)) &  | 
|
2342  | 
(\<forall>Set2 Intersection Element Set1. intersection(Set1::'a,Set2,Intersection) & member(Element::'a,Intersection) --> member(Element::'a,Set1)) &  | 
|
2343  | 
(\<forall>Set1 Intersection Element Set2. intersection(Set1::'a,Set2,Intersection) & member(Element::'a,Intersection) --> member(Element::'a,Set2)) &  | 
|
2344  | 
(\<forall>Set2 Set1 Element Intersection. intersection(Set1::'a,Set2,Intersection) & member(Element::'a,Set2) & member(Element::'a,Set1) --> member(Element::'a,Intersection)) &  | 
|
2345  | 
(\<forall>Set2 Intersection Set1. member(h(Set1::'a,Set2,Intersection),Intersection) | intersection(Set1::'a,Set2,Intersection) | member(h(Set1::'a,Set2,Intersection),Set1)) &  | 
|
2346  | 
(\<forall>Set1 Intersection Set2. member(h(Set1::'a,Set2,Intersection),Intersection) | intersection(Set1::'a,Set2,Intersection) | member(h(Set1::'a,Set2,Intersection),Set2)) &  | 
|
2347  | 
(\<forall>Set1 Set2 Intersection. member(h(Set1::'a,Set2,Intersection),Intersection) & member(h(Set1::'a,Set2,Intersection),Set2) & member(h(Set1::'a,Set2,Intersection),Set1) --> intersection(Set1::'a,Set2,Intersection)) &  | 
|
2348  | 
(intersection(a::'a,b,aIb)) &  | 
|
2349  | 
(intersection(b::'a,c,bIc)) &  | 
|
2350  | 
(intersection(a::'a,bIc,aIbIc)) &  | 
|
| 24127 | 2351  | 
(~intersection(aIb::'a,c,aIbIc)) --> False"  | 
2352  | 
oops  | 
|
2353  | 
||
2354  | 
||
2355  | 
(*6450 inferences so far. Searching to depth 14. 4.2 secs*)  | 
|
2356  | 
lemma SET009_1:  | 
|
| 24128 | 2357  | 
"(\<forall>Subset Element Superset. member(Element::'a,Subset) & ssubset(Subset::'a,Superset) --> member(Element::'a,Superset)) &  | 
2358  | 
(\<forall>Superset Subset. ssubset(Subset::'a,Superset) | member(member_of_1_not_of_2(Subset::'a,Superset),Subset)) &  | 
|
2359  | 
(\<forall>Subset Superset. member(member_of_1_not_of_2(Subset::'a,Superset),Superset) --> ssubset(Subset::'a,Superset)) &  | 
|
2360  | 
(\<forall>Subset Superset. equal_sets(Subset::'a,Superset) --> ssubset(Subset::'a,Superset)) &  | 
|
2361  | 
(\<forall>Subset Superset. equal_sets(Superset::'a,Subset) --> ssubset(Subset::'a,Superset)) &  | 
|
2362  | 
(\<forall>Set2 Set1. ssubset(Set1::'a,Set2) & ssubset(Set2::'a,Set1) --> equal_sets(Set2::'a,Set1)) &  | 
|
2363  | 
(\<forall>Set2 Difference Element Set1. difference(Set1::'a,Set2,Difference) & member(Element::'a,Difference) --> member(Element::'a,Set1)) &  | 
|
2364  | 
(\<forall>Element A_set Set1 Set2. ~(member(Element::'a,Set1) & member(Element::'a,Set2) & difference(A_set::'a,Set1,Set2))) &  | 
|
2365  | 
(\<forall>Set1 Difference Element Set2. member(Element::'a,Set1) & difference(Set1::'a,Set2,Difference) --> member(Element::'a,Difference) | member(Element::'a,Set2)) &  | 
|
2366  | 
(\<forall>Set1 Set2 Difference. difference(Set1::'a,Set2,Difference) | member(k(Set1::'a,Set2,Difference),Set1) | member(k(Set1::'a,Set2,Difference),Difference)) &  | 
|
2367  | 
(\<forall>Set1 Set2 Difference. member(k(Set1::'a,Set2,Difference),Set2) --> member(k(Set1::'a,Set2,Difference),Difference) | difference(Set1::'a,Set2,Difference)) &  | 
|
2368  | 
(\<forall>Set1 Set2 Difference. member(k(Set1::'a,Set2,Difference),Difference) & member(k(Set1::'a,Set2,Difference),Set1) --> member(k(Set1::'a,Set2,Difference),Set2) | difference(Set1::'a,Set2,Difference)) &  | 
|
2369  | 
(ssubset(d::'a,a)) &  | 
|
2370  | 
(difference(b::'a,a,bDa)) &  | 
|
2371  | 
(difference(b::'a,d,bDd)) &  | 
|
| 24127 | 2372  | 
(~ssubset(bDa::'a,bDd)) --> False"  | 
2373  | 
by meson  | 
|
2374  | 
||
2375  | 
(*34726 inferences so far. Searching to depth 6. 2420 secs: 40 mins! BIG*)  | 
|
2376  | 
lemma SET025_4:  | 
|
| 24128 | 2377  | 
"EQU001_0_ax equal &  | 
2378  | 
(\<forall>Y X. member(X::'a,Y) --> little_set(X)) &  | 
|
2379  | 
(\<forall>X Y. little_set(f1(X::'a,Y)) | equal(X::'a,Y)) &  | 
|
2380  | 
(\<forall>X Y. member(f1(X::'a,Y),X) | member(f1(X::'a,Y),Y) | equal(X::'a,Y)) &  | 
|
2381  | 
(\<forall>X Y. member(f1(X::'a,Y),X) & member(f1(X::'a,Y),Y) --> equal(X::'a,Y)) &  | 
|
2382  | 
(\<forall>X U Y. member(U::'a,non_ordered_pair(X::'a,Y)) --> equal(U::'a,X) | equal(U::'a,Y)) &  | 
|
2383  | 
(\<forall>Y U X. little_set(U) & equal(U::'a,X) --> member(U::'a,non_ordered_pair(X::'a,Y))) &  | 
|
2384  | 
(\<forall>X U Y. little_set(U) & equal(U::'a,Y) --> member(U::'a,non_ordered_pair(X::'a,Y))) &  | 
|
2385  | 
(\<forall>X Y. little_set(non_ordered_pair(X::'a,Y))) &  | 
|
2386  | 
(\<forall>X. equal(singleton_set(X),non_ordered_pair(X::'a,X))) &  | 
|
2387  | 
(\<forall>X Y. equal(ordered_pair(X::'a,Y),non_ordered_pair(singleton_set(X),non_ordered_pair(X::'a,Y)))) &  | 
|
2388  | 
(\<forall>X. ordered_pair_predicate(X) --> little_set(f2(X))) &  | 
|
2389  | 
(\<forall>X. ordered_pair_predicate(X) --> little_set(f3(X))) &  | 
|
2390  | 
(\<forall>X. ordered_pair_predicate(X) --> equal(X::'a,ordered_pair(f2(X),f3(X)))) &  | 
|
2391  | 
(\<forall>X Y Z. little_set(Y) & little_set(Z) & equal(X::'a,ordered_pair(Y::'a,Z)) --> ordered_pair_predicate(X)) &  | 
|
2392  | 
(\<forall>Z X. member(Z::'a,first(X)) --> little_set(f4(Z::'a,X))) &  | 
|
2393  | 
(\<forall>Z X. member(Z::'a,first(X)) --> little_set(f5(Z::'a,X))) &  | 
|
2394  | 
(\<forall>Z X. member(Z::'a,first(X)) --> equal(X::'a,ordered_pair(f4(Z::'a,X),f5(Z::'a,X)))) &  | 
|
2395  | 
(\<forall>Z X. member(Z::'a,first(X)) --> member(Z::'a,f4(Z::'a,X))) &  | 
|
2396  | 
(\<forall>X V Z U. little_set(U) & little_set(V) & equal(X::'a,ordered_pair(U::'a,V)) & member(Z::'a,U) --> member(Z::'a,first(X))) &  | 
|
2397  | 
(\<forall>Z X. member(Z::'a,second(X)) --> little_set(f6(Z::'a,X))) &  | 
|
2398  | 
(\<forall>Z X. member(Z::'a,second(X)) --> little_set(f7(Z::'a,X))) &  | 
|
2399  | 
(\<forall>Z X. member(Z::'a,second(X)) --> equal(X::'a,ordered_pair(f6(Z::'a,X),f7(Z::'a,X)))) &  | 
|
2400  | 
(\<forall>Z X. member(Z::'a,second(X)) --> member(Z::'a,f7(Z::'a,X))) &  | 
|
2401  | 
(\<forall>X U Z V. little_set(U) & little_set(V) & equal(X::'a,ordered_pair(U::'a,V)) & member(Z::'a,V) --> member(Z::'a,second(X))) &  | 
|
2402  | 
(\<forall>Z. member(Z::'a,estin) --> ordered_pair_predicate(Z)) &  | 
|
2403  | 
(\<forall>Z. member(Z::'a,estin) --> member(first(Z),second(Z))) &  | 
|
2404  | 
(\<forall>Z. little_set(Z) & ordered_pair_predicate(Z) & member(first(Z),second(Z)) --> member(Z::'a,estin)) &  | 
|
2405  | 
(\<forall>Y Z X. member(Z::'a,intersection(X::'a,Y)) --> member(Z::'a,X)) &  | 
|
2406  | 
(\<forall>X Z Y. member(Z::'a,intersection(X::'a,Y)) --> member(Z::'a,Y)) &  | 
|
2407  | 
(\<forall>X Z Y. member(Z::'a,X) & member(Z::'a,Y) --> member(Z::'a,intersection(X::'a,Y))) &  | 
|
2408  | 
(\<forall>Z X. ~(member(Z::'a,complement(X)) & member(Z::'a,X))) &  | 
|
2409  | 
(\<forall>Z X. little_set(Z) --> member(Z::'a,complement(X)) | member(Z::'a,X)) &  | 
|
2410  | 
(\<forall>X Y. equal(union(X::'a,Y),complement(intersection(complement(X),complement(Y))))) &  | 
|
2411  | 
(\<forall>Z X. member(Z::'a,domain_of(X)) --> ordered_pair_predicate(f8(Z::'a,X))) &  | 
|
2412  | 
(\<forall>Z X. member(Z::'a,domain_of(X)) --> member(f8(Z::'a,X),X)) &  | 
|
2413  | 
(\<forall>Z X. member(Z::'a,domain_of(X)) --> equal(Z::'a,first(f8(Z::'a,X)))) &  | 
|
2414  | 
(\<forall>X Z Xp. little_set(Z) & ordered_pair_predicate(Xp) & member(Xp::'a,X) & equal(Z::'a,first(Xp)) --> member(Z::'a,domain_of(X))) &  | 
|
2415  | 
(\<forall>X Y Z. member(Z::'a,cross_product(X::'a,Y)) --> ordered_pair_predicate(Z)) &  | 
|
2416  | 
(\<forall>Y Z X. member(Z::'a,cross_product(X::'a,Y)) --> member(first(Z),X)) &  | 
|
2417  | 
(\<forall>X Z Y. member(Z::'a,cross_product(X::'a,Y)) --> member(second(Z),Y)) &  | 
|
2418  | 
(\<forall>X Z Y. little_set(Z) & ordered_pair_predicate(Z) & member(first(Z),X) & member(second(Z),Y) --> member(Z::'a,cross_product(X::'a,Y))) &  | 
|
2419  | 
(\<forall>X Z. member(Z::'a,inv1 X) --> ordered_pair_predicate(Z)) &  | 
|
2420  | 
(\<forall>Z X. member(Z::'a,inv1 X) --> member(ordered_pair(second(Z),first(Z)),X)) &  | 
|
2421  | 
(\<forall>Z X. little_set(Z) & ordered_pair_predicate(Z) & member(ordered_pair(second(Z),first(Z)),X) --> member(Z::'a,inv1 X)) &  | 
|
2422  | 
(\<forall>Z X. member(Z::'a,rot_right(X)) --> little_set(f9(Z::'a,X))) &  | 
|
2423  | 
(\<forall>Z X. member(Z::'a,rot_right(X)) --> little_set(f10(Z::'a,X))) &  | 
|
2424  | 
(\<forall>Z X. member(Z::'a,rot_right(X)) --> little_set(f11(Z::'a,X))) &  | 
|
2425  | 
(\<forall>Z X. member(Z::'a,rot_right(X)) --> equal(Z::'a,ordered_pair(f9(Z::'a,X),ordered_pair(f10(Z::'a,X),f11(Z::'a,X))))) &  | 
|
2426  | 
(\<forall>Z X. member(Z::'a,rot_right(X)) --> member(ordered_pair(f10(Z::'a,X),ordered_pair(f11(Z::'a,X),f9(Z::'a,X))),X)) &  | 
|
2427  | 
(\<forall>Z V W U X. little_set(Z) & little_set(U) & little_set(V) & little_set(W) & equal(Z::'a,ordered_pair(U::'a,ordered_pair(V::'a,W))) & member(ordered_pair(V::'a,ordered_pair(W::'a,U)),X) --> member(Z::'a,rot_right(X))) &  | 
|
2428  | 
(\<forall>Z X. member(Z::'a,flip_range_of(X)) --> little_set(f12(Z::'a,X))) &  | 
|
2429  | 
(\<forall>Z X. member(Z::'a,flip_range_of(X)) --> little_set(f13(Z::'a,X))) &  | 
|
2430  | 
(\<forall>Z X. member(Z::'a,flip_range_of(X)) --> little_set(f14(Z::'a,X))) &  | 
|
2431  | 
(\<forall>Z X. member(Z::'a,flip_range_of(X)) --> equal(Z::'a,ordered_pair(f12(Z::'a,X),ordered_pair(f13(Z::'a,X),f14(Z::'a,X))))) &  | 
|
2432  | 
(\<forall>Z X. member(Z::'a,flip_range_of(X)) --> member(ordered_pair(f12(Z::'a,X),ordered_pair(f14(Z::'a,X),f13(Z::'a,X))),X)) &  | 
|
2433  | 
(\<forall>Z U W V X. little_set(Z) & little_set(U) & little_set(V) & little_set(W) & equal(Z::'a,ordered_pair(U::'a,ordered_pair(V::'a,W))) & member(ordered_pair(U::'a,ordered_pair(W::'a,V)),X) --> member(Z::'a,flip_range_of(X))) &  | 
|
2434  | 
(\<forall>X. equal(successor(X),union(X::'a,singleton_set(X)))) &  | 
|
2435  | 
(\<forall>Z. ~member(Z::'a,empty_set)) &  | 
|
2436  | 
(\<forall>Z. little_set(Z) --> member(Z::'a,universal_set)) &  | 
|
2437  | 
(little_set(infinity)) &  | 
|
2438  | 
(member(empty_set::'a,infinity)) &  | 
|
2439  | 
(\<forall>X. member(X::'a,infinity) --> member(successor(X),infinity)) &  | 
|
2440  | 
(\<forall>Z X. member(Z::'a,sigma(X)) --> member(f16(Z::'a,X),X)) &  | 
|
2441  | 
(\<forall>Z X. member(Z::'a,sigma(X)) --> member(Z::'a,f16(Z::'a,X))) &  | 
|
2442  | 
(\<forall>X Z Y. member(Y::'a,X) & member(Z::'a,Y) --> member(Z::'a,sigma(X))) &  | 
|
2443  | 
(\<forall>U. little_set(U) --> little_set(sigma(U))) &  | 
|
2444  | 
(\<forall>X U Y. ssubset(X::'a,Y) & member(U::'a,X) --> member(U::'a,Y)) &  | 
|
2445  | 
(\<forall>Y X. ssubset(X::'a,Y) | member(f17(X::'a,Y),X)) &  | 
|
2446  | 
(\<forall>X Y. member(f17(X::'a,Y),Y) --> ssubset(X::'a,Y)) &  | 
|
2447  | 
(\<forall>X Y. proper_subset(X::'a,Y) --> ssubset(X::'a,Y)) &  | 
|
2448  | 
(\<forall>X Y. ~(proper_subset(X::'a,Y) & equal(X::'a,Y))) &  | 
|
2449  | 
(\<forall>X Y. ssubset(X::'a,Y) --> proper_subset(X::'a,Y) | equal(X::'a,Y)) &  | 
|
2450  | 
(\<forall>Z X. member(Z::'a,powerset(X)) --> ssubset(Z::'a,X)) &  | 
|
2451  | 
(\<forall>Z X. little_set(Z) & ssubset(Z::'a,X) --> member(Z::'a,powerset(X))) &  | 
|
2452  | 
(\<forall>U. little_set(U) --> little_set(powerset(U))) &  | 
|
2453  | 
(\<forall>Z X. relation(Z) & member(X::'a,Z) --> ordered_pair_predicate(X)) &  | 
|
2454  | 
(\<forall>Z. relation(Z) | member(f18(Z),Z)) &  | 
|
2455  | 
(\<forall>Z. ordered_pair_predicate(f18(Z)) --> relation(Z)) &  | 
|
2456  | 
(\<forall>U X V W. single_valued_set(X) & little_set(U) & little_set(V) & little_set(W) & member(ordered_pair(U::'a,V),X) & member(ordered_pair(U::'a,W),X) --> equal(V::'a,W)) &  | 
|
2457  | 
(\<forall>X. single_valued_set(X) | little_set(f19(X))) &  | 
|
2458  | 
(\<forall>X. single_valued_set(X) | little_set(f20(X))) &  | 
|
2459  | 
(\<forall>X. single_valued_set(X) | little_set(f21(X))) &  | 
|
2460  | 
(\<forall>X. single_valued_set(X) | member(ordered_pair(f19(X),f20(X)),X)) &  | 
|
2461  | 
(\<forall>X. single_valued_set(X) | member(ordered_pair(f19(X),f21(X)),X)) &  | 
|
2462  | 
(\<forall>X. equal(f20(X),f21(X)) --> single_valued_set(X)) &  | 
|
2463  | 
(\<forall>Xf. function(Xf) --> relation(Xf)) &  | 
|
2464  | 
(\<forall>Xf. function(Xf) --> single_valued_set(Xf)) &  | 
|
2465  | 
(\<forall>Xf. relation(Xf) & single_valued_set(Xf) --> function(Xf)) &  | 
|
2466  | 
(\<forall>Z X Xf. member(Z::'a,image'(X::'a,Xf)) --> ordered_pair_predicate(f22(Z::'a,X,Xf))) &  | 
|
2467  | 
(\<forall>Z X Xf. member(Z::'a,image'(X::'a,Xf)) --> member(f22(Z::'a,X,Xf),Xf)) &  | 
|
2468  | 
(\<forall>Z Xf X. member(Z::'a,image'(X::'a,Xf)) --> member(first(f22(Z::'a,X,Xf)),X)) &  | 
|
2469  | 
(\<forall>X Xf Z. member(Z::'a,image'(X::'a,Xf)) --> equal(second(f22(Z::'a,X,Xf)),Z)) &  | 
|
2470  | 
(\<forall>Xf X Y Z. little_set(Z) & ordered_pair_predicate(Y) & member(Y::'a,Xf) & member(first(Y),X) & equal(second(Y),Z) --> member(Z::'a,image'(X::'a,Xf))) &  | 
|
2471  | 
(\<forall>X Xf. little_set(X) & function(Xf) --> little_set(image'(X::'a,Xf))) &  | 
|
2472  | 
(\<forall>X U Y. ~(disjoint(X::'a,Y) & member(U::'a,X) & member(U::'a,Y))) &  | 
|
2473  | 
(\<forall>Y X. disjoint(X::'a,Y) | member(f23(X::'a,Y),X)) &  | 
|
2474  | 
(\<forall>X Y. disjoint(X::'a,Y) | member(f23(X::'a,Y),Y)) &  | 
|
2475  | 
(\<forall>X. equal(X::'a,empty_set) | member(f24(X),X)) &  | 
|
2476  | 
(\<forall>X. equal(X::'a,empty_set) | disjoint(f24(X),X)) &  | 
|
2477  | 
(function(f25)) &  | 
|
2478  | 
(\<forall>X. little_set(X) --> equal(X::'a,empty_set) | member(f26(X),X)) &  | 
|
2479  | 
(\<forall>X. little_set(X) --> equal(X::'a,empty_set) | member(ordered_pair(X::'a,f26(X)),f25)) &  | 
|
2480  | 
(\<forall>Z X. member(Z::'a,range_of(X)) --> ordered_pair_predicate(f27(Z::'a,X))) &  | 
|
2481  | 
(\<forall>Z X. member(Z::'a,range_of(X)) --> member(f27(Z::'a,X),X)) &  | 
|
2482  | 
(\<forall>Z X. member(Z::'a,range_of(X)) --> equal(Z::'a,second(f27(Z::'a,X)))) &  | 
|
2483  | 
(\<forall>X Z Xp. little_set(Z) & ordered_pair_predicate(Xp) & member(Xp::'a,X) & equal(Z::'a,second(Xp)) --> member(Z::'a,range_of(X))) &  | 
|
2484  | 
(\<forall>Z. member(Z::'a,identity_relation) --> ordered_pair_predicate(Z)) &  | 
|
2485  | 
(\<forall>Z. member(Z::'a,identity_relation) --> equal(first(Z),second(Z))) &  | 
|
2486  | 
(\<forall>Z. little_set(Z) & ordered_pair_predicate(Z) & equal(first(Z),second(Z)) --> member(Z::'a,identity_relation)) &  | 
|
2487  | 
(\<forall>X Y. equal(restrct(X::'a,Y),intersection(X::'a,cross_product(Y::'a,universal_set)))) &  | 
|
2488  | 
(\<forall>Xf. one_to_one_function(Xf) --> function(Xf)) &  | 
|
2489  | 
(\<forall>Xf. one_to_one_function(Xf) --> function(inv1 Xf)) &  | 
|
2490  | 
(\<forall>Xf. function(Xf) & function(inv1 Xf) --> one_to_one_function(Xf)) &  | 
|
2491  | 
(\<forall>Z Xf Y. member(Z::'a,apply(Xf::'a,Y)) --> ordered_pair_predicate(f28(Z::'a,Xf,Y))) &  | 
|
2492  | 
(\<forall>Z Y Xf. member(Z::'a,apply(Xf::'a,Y)) --> member(f28(Z::'a,Xf,Y),Xf)) &  | 
|
2493  | 
(\<forall>Z Xf Y. member(Z::'a,apply(Xf::'a,Y)) --> equal(first(f28(Z::'a,Xf,Y)),Y)) &  | 
|
2494  | 
(\<forall>Z Xf Y. member(Z::'a,apply(Xf::'a,Y)) --> member(Z::'a,second(f28(Z::'a,Xf,Y)))) &  | 
|
2495  | 
(\<forall>Xf Y Z W. ordered_pair_predicate(W) & member(W::'a,Xf) & equal(first(W),Y) & member(Z::'a,second(W)) --> member(Z::'a,apply(Xf::'a,Y))) &  | 
|
2496  | 
(\<forall>Xf X Y. equal(apply_to_two_arguments(Xf::'a,X,Y),apply(Xf::'a,ordered_pair(X::'a,Y)))) &  | 
|
2497  | 
(\<forall>X Y Xf. maps(Xf::'a,X,Y) --> function(Xf)) &  | 
|
2498  | 
(\<forall>Y Xf X. maps(Xf::'a,X,Y) --> equal(domain_of(Xf),X)) &  | 
|
2499  | 
(\<forall>X Xf Y. maps(Xf::'a,X,Y) --> ssubset(range_of(Xf),Y)) &  | 
|
2500  | 
(\<forall>X Xf Y. function(Xf) & equal(domain_of(Xf),X) & ssubset(range_of(Xf),Y) --> maps(Xf::'a,X,Y)) &  | 
|
2501  | 
(\<forall>Xf Xs. closed(Xs::'a,Xf) --> little_set(Xs)) &  | 
|
2502  | 
(\<forall>Xs Xf. closed(Xs::'a,Xf) --> little_set(Xf)) &  | 
|
2503  | 
(\<forall>Xf Xs. closed(Xs::'a,Xf) --> maps(Xf::'a,cross_product(Xs::'a,Xs),Xs)) &  | 
|
2504  | 
(\<forall>Xf Xs. little_set(Xs) & little_set(Xf) & maps(Xf::'a,cross_product(Xs::'a,Xs),Xs) --> closed(Xs::'a,Xf)) &  | 
|
2505  | 
(\<forall>Z Xf Xg. member(Z::'a,composition(Xf::'a,Xg)) --> little_set(f29(Z::'a,Xf,Xg))) &  | 
|
2506  | 
(\<forall>Z Xf Xg. member(Z::'a,composition(Xf::'a,Xg)) --> little_set(f30(Z::'a,Xf,Xg))) &  | 
|
2507  | 
(\<forall>Z Xf Xg. member(Z::'a,composition(Xf::'a,Xg)) --> little_set(f31(Z::'a,Xf,Xg))) &  | 
|
2508  | 
(\<forall>Z Xf Xg. member(Z::'a,composition(Xf::'a,Xg)) --> equal(Z::'a,ordered_pair(f29(Z::'a,Xf,Xg),f30(Z::'a,Xf,Xg)))) &  | 
|
2509  | 
(\<forall>Z Xg Xf. member(Z::'a,composition(Xf::'a,Xg)) --> member(ordered_pair(f29(Z::'a,Xf,Xg),f31(Z::'a,Xf,Xg)),Xf)) &  | 
|
2510  | 
(\<forall>Z Xf Xg. member(Z::'a,composition(Xf::'a,Xg)) --> member(ordered_pair(f31(Z::'a,Xf,Xg),f30(Z::'a,Xf,Xg)),Xg)) &  | 
|
2511  | 
(\<forall>Z X Xf W Y Xg. little_set(Z) & little_set(X) & little_set(Y) & little_set(W) & equal(Z::'a,ordered_pair(X::'a,Y)) & member(ordered_pair(X::'a,W),Xf) & member(ordered_pair(W::'a,Y),Xg) --> member(Z::'a,composition(Xf::'a,Xg))) &  | 
|
2512  | 
(\<forall>Xh Xs2 Xf2 Xs1 Xf1. homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2) --> closed(Xs1::'a,Xf1)) &  | 
|
2513  | 
(\<forall>Xh Xs1 Xf1 Xs2 Xf2. homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2) --> closed(Xs2::'a,Xf2)) &  | 
|
2514  | 
(\<forall>Xf1 Xf2 Xh Xs1 Xs2. homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2) --> maps(Xh::'a,Xs1,Xs2)) &  | 
|
2515  | 
(\<forall>Xs2 Xs1 Xf1 Xf2 X Xh Y. homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2) & member(X::'a,Xs1) & member(Y::'a,Xs1) --> equal(apply(Xh::'a,apply_to_two_arguments(Xf1::'a,X,Y)),apply_to_two_arguments(Xf2::'a,apply(Xh::'a,X),apply(Xh::'a,Y)))) &  | 
|
2516  | 
(\<forall>Xh Xf1 Xs2 Xf2 Xs1. closed(Xs1::'a,Xf1) & closed(Xs2::'a,Xf2) & maps(Xh::'a,Xs1,Xs2) --> homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2) | member(f32(Xh::'a,Xs1,Xf1,Xs2,Xf2),Xs1)) &  | 
|
2517  | 
(\<forall>Xh Xf1 Xs2 Xf2 Xs1. closed(Xs1::'a,Xf1) & closed(Xs2::'a,Xf2) & maps(Xh::'a,Xs1,Xs2) --> homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2) | member(f33(Xh::'a,Xs1,Xf1,Xs2,Xf2),Xs1)) &  | 
|
2518  | 
(\<forall>Xh Xs1 Xf1 Xs2 Xf2. closed(Xs1::'a,Xf1) & closed(Xs2::'a,Xf2) & maps(Xh::'a,Xs1,Xs2) & equal(apply(Xh::'a,apply_to_two_arguments(Xf1::'a,f32(Xh::'a,Xs1,Xf1,Xs2,Xf2),f33(Xh::'a,Xs1,Xf1,Xs2,Xf2))),apply_to_two_arguments(Xf2::'a,apply(Xh::'a,f32(Xh::'a,Xs1,Xf1,Xs2,Xf2)),apply(Xh::'a,f33(Xh::'a,Xs1,Xf1,Xs2,Xf2)))) --> homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2)) &  | 
|
2519  | 
(\<forall>A B C. equal(A::'a,B) --> equal(f1(A::'a,C),f1(B::'a,C))) &  | 
|
2520  | 
(\<forall>D F' E. equal(D::'a,E) --> equal(f1(F'::'a,D),f1(F'::'a,E))) &  | 
|
2521  | 
(\<forall>A2 B2. equal(A2::'a,B2) --> equal(f2(A2),f2(B2))) &  | 
|
2522  | 
(\<forall>G4 H4. equal(G4::'a,H4) --> equal(f3(G4),f3(H4))) &  | 
|
2523  | 
(\<forall>O7 P7 Q7. equal(O7::'a,P7) --> equal(f4(O7::'a,Q7),f4(P7::'a,Q7))) &  | 
|
2524  | 
(\<forall>R7 T7 S7. equal(R7::'a,S7) --> equal(f4(T7::'a,R7),f4(T7::'a,S7))) &  | 
|
2525  | 
(\<forall>U7 V7 W7. equal(U7::'a,V7) --> equal(f5(U7::'a,W7),f5(V7::'a,W7))) &  | 
|
2526  | 
(\<forall>X7 Z7 Y7. equal(X7::'a,Y7) --> equal(f5(Z7::'a,X7),f5(Z7::'a,Y7))) &  | 
|
2527  | 
(\<forall>A8 B8 C8. equal(A8::'a,B8) --> equal(f6(A8::'a,C8),f6(B8::'a,C8))) &  | 
|
2528  | 
(\<forall>D8 F8 E8. equal(D8::'a,E8) --> equal(f6(F8::'a,D8),f6(F8::'a,E8))) &  | 
|
2529  | 
(\<forall>G8 H8 I8. equal(G8::'a,H8) --> equal(f7(G8::'a,I8),f7(H8::'a,I8))) &  | 
|
2530  | 
(\<forall>J8 L8 K8. equal(J8::'a,K8) --> equal(f7(L8::'a,J8),f7(L8::'a,K8))) &  | 
|
2531  | 
(\<forall>M8 N8 O8. equal(M8::'a,N8) --> equal(f8(M8::'a,O8),f8(N8::'a,O8))) &  | 
|
2532  | 
(\<forall>P8 R8 Q8. equal(P8::'a,Q8) --> equal(f8(R8::'a,P8),f8(R8::'a,Q8))) &  | 
|
2533  | 
(\<forall>S8 T8 U8. equal(S8::'a,T8) --> equal(f9(S8::'a,U8),f9(T8::'a,U8))) &  | 
|
2534  | 
(\<forall>V8 X8 W8. equal(V8::'a,W8) --> equal(f9(X8::'a,V8),f9(X8::'a,W8))) &  | 
|
2535  | 
(\<forall>G H I'. equal(G::'a,H) --> equal(f10(G::'a,I'),f10(H::'a,I'))) &  | 
|
2536  | 
(\<forall>J L K'. equal(J::'a,K') --> equal(f10(L::'a,J),f10(L::'a,K'))) &  | 
|
2537  | 
(\<forall>M N O'. equal(M::'a,N) --> equal(f11(M::'a,O'),f11(N::'a,O'))) &  | 
|
2538  | 
(\<forall>P R Q. equal(P::'a,Q) --> equal(f11(R::'a,P),f11(R::'a,Q))) &  | 
|
2539  | 
(\<forall>S' T' U. equal(S'::'a,T') --> equal(f12(S'::'a,U),f12(T'::'a,U))) &  | 
|
2540  | 
(\<forall>V X W. equal(V::'a,W) --> equal(f12(X::'a,V),f12(X::'a,W))) &  | 
|
2541  | 
(\<forall>Y Z A1. equal(Y::'a,Z) --> equal(f13(Y::'a,A1),f13(Z::'a,A1))) &  | 
|
2542  | 
(\<forall>B1 D1 C1. equal(B1::'a,C1) --> equal(f13(D1::'a,B1),f13(D1::'a,C1))) &  | 
|
2543  | 
(\<forall>E1 F1 G1. equal(E1::'a,F1) --> equal(f14(E1::'a,G1),f14(F1::'a,G1))) &  | 
|
2544  | 
(\<forall>H1 J1 I1. equal(H1::'a,I1) --> equal(f14(J1::'a,H1),f14(J1::'a,I1))) &  | 
|
2545  | 
(\<forall>K1 L1 M1. equal(K1::'a,L1) --> equal(f16(K1::'a,M1),f16(L1::'a,M1))) &  | 
|
2546  | 
(\<forall>N1 P1 O1. equal(N1::'a,O1) --> equal(f16(P1::'a,N1),f16(P1::'a,O1))) &  | 
|
2547  | 
(\<forall>Q1 R1 S1. equal(Q1::'a,R1) --> equal(f17(Q1::'a,S1),f17(R1::'a,S1))) &  | 
|
2548  | 
(\<forall>T1 V1 U1. equal(T1::'a,U1) --> equal(f17(V1::'a,T1),f17(V1::'a,U1))) &  | 
|
2549  | 
(\<forall>W1 X1. equal(W1::'a,X1) --> equal(f18(W1),f18(X1))) &  | 
|
2550  | 
(\<forall>Y1 Z1. equal(Y1::'a,Z1) --> equal(f19(Y1),f19(Z1))) &  | 
|
2551  | 
(\<forall>C2 D2. equal(C2::'a,D2) --> equal(f20(C2),f20(D2))) &  | 
|
2552  | 
(\<forall>E2 F2. equal(E2::'a,F2) --> equal(f21(E2),f21(F2))) &  | 
|
2553  | 
(\<forall>G2 H2 I2 J2. equal(G2::'a,H2) --> equal(f22(G2::'a,I2,J2),f22(H2::'a,I2,J2))) &  | 
|
2554  | 
(\<forall>K2 M2 L2 N2. equal(K2::'a,L2) --> equal(f22(M2::'a,K2,N2),f22(M2::'a,L2,N2))) &  | 
|
2555  | 
(\<forall>O2 Q2 R2 P2. equal(O2::'a,P2) --> equal(f22(Q2::'a,R2,O2),f22(Q2::'a,R2,P2))) &  | 
|
2556  | 
(\<forall>S2 T2 U2. equal(S2::'a,T2) --> equal(f23(S2::'a,U2),f23(T2::'a,U2))) &  | 
|
2557  | 
(\<forall>V2 X2 W2. equal(V2::'a,W2) --> equal(f23(X2::'a,V2),f23(X2::'a,W2))) &  | 
|
2558  | 
(\<forall>Y2 Z2. equal(Y2::'a,Z2) --> equal(f24(Y2),f24(Z2))) &  | 
|
2559  | 
(\<forall>A3 B3. equal(A3::'a,B3) --> equal(f26(A3),f26(B3))) &  | 
|
2560  | 
(\<forall>C3 D3 E3. equal(C3::'a,D3) --> equal(f27(C3::'a,E3),f27(D3::'a,E3))) &  | 
|
2561  | 
(\<forall>F3 H3 G3. equal(F3::'a,G3) --> equal(f27(H3::'a,F3),f27(H3::'a,G3))) &  | 
|
2562  | 
(\<forall>I3 J3 K3 L3. equal(I3::'a,J3) --> equal(f28(I3::'a,K3,L3),f28(J3::'a,K3,L3))) &  | 
|
2563  | 
(\<forall>M3 O3 N3 P3. equal(M3::'a,N3) --> equal(f28(O3::'a,M3,P3),f28(O3::'a,N3,P3))) &  | 
|
2564  | 
(\<forall>Q3 S3 T3 R3. equal(Q3::'a,R3) --> equal(f28(S3::'a,T3,Q3),f28(S3::'a,T3,R3))) &  | 
|
2565  | 
(\<forall>U3 V3 W3 X3. equal(U3::'a,V3) --> equal(f29(U3::'a,W3,X3),f29(V3::'a,W3,X3))) &  | 
|
2566  | 
(\<forall>Y3 A4 Z3 B4. equal(Y3::'a,Z3) --> equal(f29(A4::'a,Y3,B4),f29(A4::'a,Z3,B4))) &  | 
|
2567  | 
(\<forall>C4 E4 F4 D4. equal(C4::'a,D4) --> equal(f29(E4::'a,F4,C4),f29(E4::'a,F4,D4))) &  | 
|
2568  | 
(\<forall>I4 J4 K4 L4. equal(I4::'a,J4) --> equal(f30(I4::'a,K4,L4),f30(J4::'a,K4,L4))) &  | 
|
2569  | 
(\<forall>M4 O4 N4 P4. equal(M4::'a,N4) --> equal(f30(O4::'a,M4,P4),f30(O4::'a,N4,P4))) &  | 
|
2570  | 
(\<forall>Q4 S4 T4 R4. equal(Q4::'a,R4) --> equal(f30(S4::'a,T4,Q4),f30(S4::'a,T4,R4))) &  | 
|
2571  | 
(\<forall>U4 V4 W4 X4. equal(U4::'a,V4) --> equal(f31(U4::'a,W4,X4),f31(V4::'a,W4,X4))) &  | 
|
2572  | 
(\<forall>Y4 A5 Z4 B5. equal(Y4::'a,Z4) --> equal(f31(A5::'a,Y4,B5),f31(A5::'a,Z4,B5))) &  | 
|
2573  | 
(\<forall>C5 E5 F5 D5. equal(C5::'a,D5) --> equal(f31(E5::'a,F5,C5),f31(E5::'a,F5,D5))) &  | 
|
2574  | 
(\<forall>G5 H5 I5 J5 K5 L5. equal(G5::'a,H5) --> equal(f32(G5::'a,I5,J5,K5,L5),f32(H5::'a,I5,J5,K5,L5))) &  | 
|
2575  | 
(\<forall>M5 O5 N5 P5 Q5 R5. equal(M5::'a,N5) --> equal(f32(O5::'a,M5,P5,Q5,R5),f32(O5::'a,N5,P5,Q5,R5))) &  | 
|
2576  | 
(\<forall>S5 U5 V5 T5 W5 X5. equal(S5::'a,T5) --> equal(f32(U5::'a,V5,S5,W5,X5),f32(U5::'a,V5,T5,W5,X5))) &  | 
|
2577  | 
(\<forall>Y5 A6 B6 C6 Z5 D6. equal(Y5::'a,Z5) --> equal(f32(A6::'a,B6,C6,Y5,D6),f32(A6::'a,B6,C6,Z5,D6))) &  | 
|
2578  | 
(\<forall>E6 G6 H6 I6 J6 F6. equal(E6::'a,F6) --> equal(f32(G6::'a,H6,I6,J6,E6),f32(G6::'a,H6,I6,J6,F6))) &  | 
|
2579  | 
(\<forall>K6 L6 M6 N6 O6 P6. equal(K6::'a,L6) --> equal(f33(K6::'a,M6,N6,O6,P6),f33(L6::'a,M6,N6,O6,P6))) &  | 
|
2580  | 
(\<forall>Q6 S6 R6 T6 U6 V6. equal(Q6::'a,R6) --> equal(f33(S6::'a,Q6,T6,U6,V6),f33(S6::'a,R6,T6,U6,V6))) &  | 
|
2581  | 
(\<forall>W6 Y6 Z6 X6 A7 B7. equal(W6::'a,X6) --> equal(f33(Y6::'a,Z6,W6,A7,B7),f33(Y6::'a,Z6,X6,A7,B7))) &  | 
|
2582  | 
(\<forall>C7 E7 F7 G7 D7 H7. equal(C7::'a,D7) --> equal(f33(E7::'a,F7,G7,C7,H7),f33(E7::'a,F7,G7,D7,H7))) &  | 
|
2583  | 
(\<forall>I7 K7 L7 M7 N7 J7. equal(I7::'a,J7) --> equal(f33(K7::'a,L7,M7,N7,I7),f33(K7::'a,L7,M7,N7,J7))) &  | 
|
2584  | 
(\<forall>A B C. equal(A::'a,B) --> equal(apply(A::'a,C),apply(B::'a,C))) &  | 
|
2585  | 
(\<forall>D F' E. equal(D::'a,E) --> equal(apply(F'::'a,D),apply(F'::'a,E))) &  | 
|
2586  | 
(\<forall>G H I' J. equal(G::'a,H) --> equal(apply_to_two_arguments(G::'a,I',J),apply_to_two_arguments(H::'a,I',J))) &  | 
|
2587  | 
(\<forall>K' M L N. equal(K'::'a,L) --> equal(apply_to_two_arguments(M::'a,K',N),apply_to_two_arguments(M::'a,L,N))) &  | 
|
2588  | 
(\<forall>O' Q R P. equal(O'::'a,P) --> equal(apply_to_two_arguments(Q::'a,R,O'),apply_to_two_arguments(Q::'a,R,P))) &  | 
|
2589  | 
(\<forall>S' T'. equal(S'::'a,T') --> equal(complement(S'),complement(T'))) &  | 
|
2590  | 
(\<forall>U V W. equal(U::'a,V) --> equal(composition(U::'a,W),composition(V::'a,W))) &  | 
|
2591  | 
(\<forall>X Z Y. equal(X::'a,Y) --> equal(composition(Z::'a,X),composition(Z::'a,Y))) &  | 
|
2592  | 
(\<forall>A1 B1. equal(A1::'a,B1) --> equal(inv1 A1,inv1 B1)) &  | 
|
2593  | 
(\<forall>C1 D1 E1. equal(C1::'a,D1) --> equal(cross_product(C1::'a,E1),cross_product(D1::'a,E1))) &  | 
|
2594  | 
(\<forall>F1 H1 G1. equal(F1::'a,G1) --> equal(cross_product(H1::'a,F1),cross_product(H1::'a,G1))) &  | 
|
2595  | 
(\<forall>I1 J1. equal(I1::'a,J1) --> equal(domain_of(I1),domain_of(J1))) &  | 
|
2596  | 
(\<forall>I10 J10. equal(I10::'a,J10) --> equal(first(I10),first(J10))) &  | 
|
2597  | 
(\<forall>Q10 R10. equal(Q10::'a,R10) --> equal(flip_range_of(Q10),flip_range_of(R10))) &  | 
|
2598  | 
(\<forall>S10 T10 U10. equal(S10::'a,T10) --> equal(image'(S10::'a,U10),image'(T10::'a,U10))) &  | 
|
2599  | 
(\<forall>V10 X10 W10. equal(V10::'a,W10) --> equal(image'(X10::'a,V10),image'(X10::'a,W10))) &  | 
|
2600  | 
(\<forall>Y10 Z10 A11. equal(Y10::'a,Z10) --> equal(intersection(Y10::'a,A11),intersection(Z10::'a,A11))) &  | 
|
2601  | 
(\<forall>B11 D11 C11. equal(B11::'a,C11) --> equal(intersection(D11::'a,B11),intersection(D11::'a,C11))) &  | 
|
2602  | 
(\<forall>E11 F11 G11. equal(E11::'a,F11) --> equal(non_ordered_pair(E11::'a,G11),non_ordered_pair(F11::'a,G11))) &  | 
|
2603  | 
(\<forall>H11 J11 I11. equal(H11::'a,I11) --> equal(non_ordered_pair(J11::'a,H11),non_ordered_pair(J11::'a,I11))) &  | 
|
2604  | 
(\<forall>K11 L11 M11. equal(K11::'a,L11) --> equal(ordered_pair(K11::'a,M11),ordered_pair(L11::'a,M11))) &  | 
|
2605  | 
(\<forall>N11 P11 O11. equal(N11::'a,O11) --> equal(ordered_pair(P11::'a,N11),ordered_pair(P11::'a,O11))) &  | 
|
2606  | 
(\<forall>Q11 R11. equal(Q11::'a,R11) --> equal(powerset(Q11),powerset(R11))) &  | 
|
2607  | 
(\<forall>S11 T11. equal(S11::'a,T11) --> equal(range_of(S11),range_of(T11))) &  | 
|
2608  | 
(\<forall>U11 V11 W11. equal(U11::'a,V11) --> equal(restrct(U11::'a,W11),restrct(V11::'a,W11))) &  | 
|
2609  | 
(\<forall>X11 Z11 Y11. equal(X11::'a,Y11) --> equal(restrct(Z11::'a,X11),restrct(Z11::'a,Y11))) &  | 
|
2610  | 
(\<forall>A12 B12. equal(A12::'a,B12) --> equal(rot_right(A12),rot_right(B12))) &  | 
|
2611  | 
(\<forall>C12 D12. equal(C12::'a,D12) --> equal(second(C12),second(D12))) &  | 
|
2612  | 
(\<forall>K12 L12. equal(K12::'a,L12) --> equal(sigma(K12),sigma(L12))) &  | 
|
2613  | 
(\<forall>M12 N12. equal(M12::'a,N12) --> equal(singleton_set(M12),singleton_set(N12))) &  | 
|
2614  | 
(\<forall>O12 P12. equal(O12::'a,P12) --> equal(successor(O12),successor(P12))) &  | 
|
2615  | 
(\<forall>Q12 R12 S12. equal(Q12::'a,R12) --> equal(union(Q12::'a,S12),union(R12::'a,S12))) &  | 
|
2616  | 
(\<forall>T12 V12 U12. equal(T12::'a,U12) --> equal(union(V12::'a,T12),union(V12::'a,U12))) &  | 
|
2617  | 
(\<forall>W12 X12 Y12. equal(W12::'a,X12) & closed(W12::'a,Y12) --> closed(X12::'a,Y12)) &  | 
|
2618  | 
(\<forall>Z12 B13 A13. equal(Z12::'a,A13) & closed(B13::'a,Z12) --> closed(B13::'a,A13)) &  | 
|
2619  | 
(\<forall>C13 D13 E13. equal(C13::'a,D13) & disjoint(C13::'a,E13) --> disjoint(D13::'a,E13)) &  | 
|
2620  | 
(\<forall>F13 H13 G13. equal(F13::'a,G13) & disjoint(H13::'a,F13) --> disjoint(H13::'a,G13)) &  | 
|
2621  | 
(\<forall>I13 J13. equal(I13::'a,J13) & function(I13) --> function(J13)) &  | 
|
2622  | 
(\<forall>K13 L13 M13 N13 O13 P13. equal(K13::'a,L13) & homomorphism(K13::'a,M13,N13,O13,P13) --> homomorphism(L13::'a,M13,N13,O13,P13)) &  | 
|
2623  | 
(\<forall>Q13 S13 R13 T13 U13 V13. equal(Q13::'a,R13) & homomorphism(S13::'a,Q13,T13,U13,V13) --> homomorphism(S13::'a,R13,T13,U13,V13)) &  | 
|
2624  | 
(\<forall>W13 Y13 Z13 X13 A14 B14. equal(W13::'a,X13) & homomorphism(Y13::'a,Z13,W13,A14,B14) --> homomorphism(Y13::'a,Z13,X13,A14,B14)) &  | 
|
2625  | 
(\<forall>C14 E14 F14 G14 D14 H14. equal(C14::'a,D14) & homomorphism(E14::'a,F14,G14,C14,H14) --> homomorphism(E14::'a,F14,G14,D14,H14)) &  | 
|
2626  | 
(\<forall>I14 K14 L14 M14 N14 J14. equal(I14::'a,J14) & homomorphism(K14::'a,L14,M14,N14,I14) --> homomorphism(K14::'a,L14,M14,N14,J14)) &  | 
|
2627  | 
(\<forall>O14 P14. equal(O14::'a,P14) & little_set(O14) --> little_set(P14)) &  | 
|
2628  | 
(\<forall>Q14 R14 S14 T14. equal(Q14::'a,R14) & maps(Q14::'a,S14,T14) --> maps(R14::'a,S14,T14)) &  | 
|
2629  | 
(\<forall>U14 W14 V14 X14. equal(U14::'a,V14) & maps(W14::'a,U14,X14) --> maps(W14::'a,V14,X14)) &  | 
|
2630  | 
(\<forall>Y14 A15 B15 Z14. equal(Y14::'a,Z14) & maps(A15::'a,B15,Y14) --> maps(A15::'a,B15,Z14)) &  | 
|
2631  | 
(\<forall>C15 D15 E15. equal(C15::'a,D15) & member(C15::'a,E15) --> member(D15::'a,E15)) &  | 
|
2632  | 
(\<forall>F15 H15 G15. equal(F15::'a,G15) & member(H15::'a,F15) --> member(H15::'a,G15)) &  | 
|
2633  | 
(\<forall>I15 J15. equal(I15::'a,J15) & one_to_one_function(I15) --> one_to_one_function(J15)) &  | 
|
2634  | 
(\<forall>K15 L15. equal(K15::'a,L15) & ordered_pair_predicate(K15) --> ordered_pair_predicate(L15)) &  | 
|
2635  | 
(\<forall>M15 N15 O15. equal(M15::'a,N15) & proper_subset(M15::'a,O15) --> proper_subset(N15::'a,O15)) &  | 
|
2636  | 
(\<forall>P15 R15 Q15. equal(P15::'a,Q15) & proper_subset(R15::'a,P15) --> proper_subset(R15::'a,Q15)) &  | 
|
2637  | 
(\<forall>S15 T15. equal(S15::'a,T15) & relation(S15) --> relation(T15)) &  | 
|
2638  | 
(\<forall>U15 V15. equal(U15::'a,V15) & single_valued_set(U15) --> single_valued_set(V15)) &  | 
|
2639  | 
(\<forall>W15 X15 Y15. equal(W15::'a,X15) & ssubset(W15::'a,Y15) --> ssubset(X15::'a,Y15)) &  | 
|
2640  | 
(\<forall>Z15 B16 A16. equal(Z15::'a,A16) & ssubset(B16::'a,Z15) --> ssubset(B16::'a,A16)) &  | 
|
| 24127 | 2641  | 
(~little_set(ordered_pair(a::'a,b))) --> False"  | 
2642  | 
oops  | 
|
2643  | 
||
2644  | 
||
2645  | 
(*13 inferences so far. Searching to depth 8. 0 secs*)  | 
|
2646  | 
lemma SET046_5:  | 
|
| 24128 | 2647  | 
"(\<forall>Y X. ~(element(X::'a,a) & element(X::'a,Y) & element(Y::'a,X))) &  | 
2648  | 
(\<forall>X. element(X::'a,f(X)) | element(X::'a,a)) &  | 
|
| 24127 | 2649  | 
(\<forall>X. element(f(X),X) | element(X::'a,a)) --> False"  | 
2650  | 
by meson  | 
|
2651  | 
||
2652  | 
(*33 inferences so far. Searching to depth 9. 0.2 secs*)  | 
|
2653  | 
lemma SET047_5:  | 
|
| 24128 | 2654  | 
"(\<forall>X Z Y. set_equal(X::'a,Y) & element(Z::'a,X) --> element(Z::'a,Y)) &  | 
2655  | 
(\<forall>Y Z X. set_equal(X::'a,Y) & element(Z::'a,Y) --> element(Z::'a,X)) &  | 
|
2656  | 
(\<forall>X Y. element(f(X::'a,Y),X) | element(f(X::'a,Y),Y) | set_equal(X::'a,Y)) &  | 
|
2657  | 
(\<forall>X Y. element(f(X::'a,Y),Y) & element(f(X::'a,Y),X) --> set_equal(X::'a,Y)) &  | 
|
2658  | 
(set_equal(a::'a,b) | set_equal(b::'a,a)) &  | 
|
| 24127 | 2659  | 
(~(set_equal(b::'a,a) & set_equal(a::'a,b))) --> False"  | 
2660  | 
by meson  | 
|
2661  | 
||
2662  | 
(*311 inferences so far. Searching to depth 12. 0.1 secs*)  | 
|
2663  | 
lemma SYN034_1:  | 
|
| 24128 | 2664  | 
"(\<forall>A. p(A::'a,a) | p(A::'a,f(A))) &  | 
2665  | 
(\<forall>A. p(A::'a,a) | p(f(A),A)) &  | 
|
| 24127 | 2666  | 
(\<forall>A B. ~(p(A::'a,B) & p(B::'a,A) & p(B::'a,a))) --> False"  | 
2667  | 
by meson  | 
|
2668  | 
||
2669  | 
(*30 inferences so far. Searching to depth 6. 0.2 secs*)  | 
|
2670  | 
lemma SYN071_1:  | 
|
2671  | 
"EQU001_0_ax equal &  | 
|
| 24128 | 2672  | 
(equal(a::'a,b) | equal(c::'a,d)) &  | 
2673  | 
(equal(a::'a,c) | equal(b::'a,d)) &  | 
|
2674  | 
(~equal(a::'a,d)) &  | 
|
| 24127 | 2675  | 
(~equal(b::'a,c)) --> False"  | 
2676  | 
by meson  | 
|
2677  | 
||
2678  | 
(*1897410 inferences so far. Searching to depth 48  | 
|
2679  | 
206s, nearly 4 mins on griffon.*)  | 
|
2680  | 
lemma SYN349_1:  | 
|
| 24128 | 2681  | 
"(\<forall>X Y. f(w(X),g(X::'a,Y)) --> f(X::'a,g(X::'a,Y))) &  | 
2682  | 
(\<forall>X Y. f(X::'a,g(X::'a,Y)) --> f(w(X),g(X::'a,Y))) &  | 
|
2683  | 
(\<forall>Y X. f(X::'a,g(X::'a,Y)) & f(Y::'a,g(X::'a,Y)) --> f(g(X::'a,Y),Y) | f(g(X::'a,Y),w(X))) &  | 
|
2684  | 
(\<forall>Y X. f(g(X::'a,Y),Y) & f(Y::'a,g(X::'a,Y)) --> f(X::'a,g(X::'a,Y)) | f(g(X::'a,Y),w(X))) &  | 
|
2685  | 
(\<forall>Y X. f(X::'a,g(X::'a,Y)) | f(g(X::'a,Y),Y) | f(Y::'a,g(X::'a,Y)) | f(g(X::'a,Y),w(X))) &  | 
|
2686  | 
(\<forall>Y X. f(X::'a,g(X::'a,Y)) & f(g(X::'a,Y),Y) --> f(Y::'a,g(X::'a,Y)) | f(g(X::'a,Y),w(X))) &  | 
|
2687  | 
(\<forall>Y X. f(X::'a,g(X::'a,Y)) & f(g(X::'a,Y),w(X)) --> f(g(X::'a,Y),Y) | f(Y::'a,g(X::'a,Y))) &  | 
|
2688  | 
(\<forall>Y X. f(g(X::'a,Y),Y) & f(g(X::'a,Y),w(X)) --> f(X::'a,g(X::'a,Y)) | f(Y::'a,g(X::'a,Y))) &  | 
|
2689  | 
(\<forall>Y X. f(Y::'a,g(X::'a,Y)) & f(g(X::'a,Y),w(X)) --> f(X::'a,g(X::'a,Y)) | f(g(X::'a,Y),Y)) &  | 
|
| 24127 | 2690  | 
(\<forall>Y X. ~(f(X::'a,g(X::'a,Y)) & f(g(X::'a,Y),Y) & f(Y::'a,g(X::'a,Y)) & f(g(X::'a,Y),w(X)))) --> False"  | 
2691  | 
oops  | 
|
2692  | 
||
2693  | 
(*398 inferences so far. Searching to depth 12. 0.4 secs*)  | 
|
2694  | 
lemma SYN352_1:  | 
|
| 24128 | 2695  | 
"(f(a::'a,b)) &  | 
2696  | 
(\<forall>X Y. f(X::'a,Y) --> f(b::'a,z(X::'a,Y)) | f(Y::'a,z(X::'a,Y))) &  | 
|
2697  | 
(\<forall>X Y. f(X::'a,Y) | f(z(X::'a,Y),z(X::'a,Y))) &  | 
|
2698  | 
(\<forall>X Y. f(b::'a,z(X::'a,Y)) | f(X::'a,z(X::'a,Y)) | f(z(X::'a,Y),z(X::'a,Y))) &  | 
|
2699  | 
(\<forall>X Y. f(b::'a,z(X::'a,Y)) & f(X::'a,z(X::'a,Y)) --> f(z(X::'a,Y),z(X::'a,Y))) &  | 
|
2700  | 
(\<forall>X Y. ~(f(X::'a,Y) & f(X::'a,z(X::'a,Y)) & f(Y::'a,z(X::'a,Y)))) &  | 
|
| 24127 | 2701  | 
(\<forall>X Y. f(X::'a,Y) --> f(X::'a,z(X::'a,Y)) | f(Y::'a,z(X::'a,Y))) --> False"  | 
2702  | 
by meson  | 
|
2703  | 
||
2704  | 
(*5336 inferences so far. Searching to depth 15. 5.3 secs*)  | 
|
2705  | 
lemma TOP001_2:  | 
|
| 24128 | 2706  | 
"(\<forall>Vf U. element_of_set(U::'a,union_of_members(Vf)) --> element_of_set(U::'a,f1(Vf::'a,U))) &  | 
2707  | 
(\<forall>U Vf. element_of_set(U::'a,union_of_members(Vf)) --> element_of_collection(f1(Vf::'a,U),Vf)) &  | 
|
2708  | 
(\<forall>U Uu1 Vf. element_of_set(U::'a,Uu1) & element_of_collection(Uu1::'a,Vf) --> element_of_set(U::'a,union_of_members(Vf))) &  | 
|
2709  | 
(\<forall>Vf X. basis(X::'a,Vf) --> equal_sets(union_of_members(Vf),X)) &  | 
|
2710  | 
(\<forall>Vf U X. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> element_of_set(X::'a,f10(Vf::'a,U,X))) &  | 
|
2711  | 
(\<forall>U X Vf. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> element_of_collection(f10(Vf::'a,U,X),Vf)) &  | 
|
2712  | 
(\<forall>X. subset_sets(X::'a,X)) &  | 
|
2713  | 
(\<forall>X U Y. subset_sets(X::'a,Y) & element_of_set(U::'a,X) --> element_of_set(U::'a,Y)) &  | 
|
2714  | 
(\<forall>X Y. equal_sets(X::'a,Y) --> subset_sets(X::'a,Y)) &  | 
|
2715  | 
(\<forall>Y X. subset_sets(X::'a,Y) | element_of_set(in_1st_set(X::'a,Y),X)) &  | 
|
2716  | 
(\<forall>X Y. element_of_set(in_1st_set(X::'a,Y),Y) --> subset_sets(X::'a,Y)) &  | 
|
2717  | 
(basis(cx::'a,f)) &  | 
|
| 24127 | 2718  | 
(~subset_sets(union_of_members(top_of_basis(f)),cx)) --> False"  | 
2719  | 
by meson  | 
|
2720  | 
||
2721  | 
(*0 inferences so far. Searching to depth 0. 0 secs*)  | 
|
2722  | 
lemma TOP002_2:  | 
|
| 24128 | 2723  | 
"(\<forall>Vf U. element_of_collection(U::'a,top_of_basis(Vf)) | element_of_set(f11(Vf::'a,U),U)) &  | 
2724  | 
(\<forall>X. ~element_of_set(X::'a,empty_set)) &  | 
|
| 24127 | 2725  | 
(~element_of_collection(empty_set::'a,top_of_basis(f))) --> False"  | 
2726  | 
by meson  | 
|
2727  | 
||
2728  | 
(*0 inferences so far. Searching to depth 0. 6.5 secs. BIG*)  | 
|
2729  | 
lemma TOP004_1:  | 
|
| 24128 | 2730  | 
"(\<forall>Vf U. element_of_set(U::'a,union_of_members(Vf)) --> element_of_set(U::'a,f1(Vf::'a,U))) &  | 
2731  | 
(\<forall>U Vf. element_of_set(U::'a,union_of_members(Vf)) --> element_of_collection(f1(Vf::'a,U),Vf)) &  | 
|
2732  | 
(\<forall>U Uu1 Vf. element_of_set(U::'a,Uu1) & element_of_collection(Uu1::'a,Vf) --> element_of_set(U::'a,union_of_members(Vf))) &  | 
|
2733  | 
(\<forall>Vf U Va. element_of_set(U::'a,intersection_of_members(Vf)) & element_of_collection(Va::'a,Vf) --> element_of_set(U::'a,Va)) &  | 
|
2734  | 
(\<forall>U Vf. element_of_set(U::'a,intersection_of_members(Vf)) | element_of_collection(f2(Vf::'a,U),Vf)) &  | 
|
2735  | 
(\<forall>Vf U. element_of_set(U::'a,f2(Vf::'a,U)) --> element_of_set(U::'a,intersection_of_members(Vf))) &  | 
|
2736  | 
(\<forall>Vt X. topological_space(X::'a,Vt) --> equal_sets(union_of_members(Vt),X)) &  | 
|
2737  | 
(\<forall>X Vt. topological_space(X::'a,Vt) --> element_of_collection(empty_set::'a,Vt)) &  | 
|
2738  | 
(\<forall>X Vt. topological_space(X::'a,Vt) --> element_of_collection(X::'a,Vt)) &  | 
|
2739  | 
(\<forall>X Y Z Vt. topological_space(X::'a,Vt) & element_of_collection(Y::'a,Vt) & element_of_collection(Z::'a,Vt) --> element_of_collection(intersection_of_sets(Y::'a,Z),Vt)) &  | 
|
2740  | 
(\<forall>X Vf Vt. topological_space(X::'a,Vt) & subset_collections(Vf::'a,Vt) --> element_of_collection(union_of_members(Vf),Vt)) &  | 
|
2741  | 
(\<forall>X Vt. equal_sets(union_of_members(Vt),X) & element_of_collection(empty_set::'a,Vt) & element_of_collection(X::'a,Vt) --> topological_space(X::'a,Vt) | element_of_collection(f3(X::'a,Vt),Vt) | subset_collections(f5(X::'a,Vt),Vt)) &  | 
|
2742  | 
(\<forall>X Vt. equal_sets(union_of_members(Vt),X) & element_of_collection(empty_set::'a,Vt) & element_of_collection(X::'a,Vt) & element_of_collection(union_of_members(f5(X::'a,Vt)),Vt) --> topological_space(X::'a,Vt) | element_of_collection(f3(X::'a,Vt),Vt)) &  | 
|
2743  | 
(\<forall>X Vt. equal_sets(union_of_members(Vt),X) & element_of_collection(empty_set::'a,Vt) & element_of_collection(X::'a,Vt) --> topological_space(X::'a,Vt) | element_of_collection(f4(X::'a,Vt),Vt) | subset_collections(f5(X::'a,Vt),Vt)) &  | 
|
2744  | 
(\<forall>X Vt. equal_sets(union_of_members(Vt),X) & element_of_collection(empty_set::'a,Vt) & element_of_collection(X::'a,Vt) & element_of_collection(union_of_members(f5(X::'a,Vt)),Vt) --> topological_space(X::'a,Vt) | element_of_collection(f4(X::'a,Vt),Vt)) &  | 
|
2745  | 
(\<forall>X Vt. equal_sets(union_of_members(Vt),X) & element_of_collection(empty_set::'a,Vt) & element_of_collection(X::'a,Vt) & element_of_collection(intersection_of_sets(f3(X::'a,Vt),f4(X::'a,Vt)),Vt) --> topological_space(X::'a,Vt) | subset_collections(f5(X::'a,Vt),Vt)) &  | 
|
2746  | 
(\<forall>X Vt. equal_sets(union_of_members(Vt),X) & element_of_collection(empty_set::'a,Vt) & element_of_collection(X::'a,Vt) & element_of_collection(intersection_of_sets(f3(X::'a,Vt),f4(X::'a,Vt)),Vt) & element_of_collection(union_of_members(f5(X::'a,Vt)),Vt) --> topological_space(X::'a,Vt)) &  | 
|
2747  | 
(\<forall>U X Vt. open(U::'a,X,Vt) --> topological_space(X::'a,Vt)) &  | 
|
2748  | 
(\<forall>X U Vt. open(U::'a,X,Vt) --> element_of_collection(U::'a,Vt)) &  | 
|
2749  | 
(\<forall>X U Vt. topological_space(X::'a,Vt) & element_of_collection(U::'a,Vt) --> open(U::'a,X,Vt)) &  | 
|
2750  | 
(\<forall>U X Vt. closed(U::'a,X,Vt) --> topological_space(X::'a,Vt)) &  | 
|
2751  | 
(\<forall>U X Vt. closed(U::'a,X,Vt) --> open(relative_complement_sets(U::'a,X),X,Vt)) &  | 
|
2752  | 
(\<forall>U X Vt. topological_space(X::'a,Vt) & open(relative_complement_sets(U::'a,X),X,Vt) --> closed(U::'a,X,Vt)) &  | 
|
2753  | 
(\<forall>Vs X Vt. finer(Vt::'a,Vs,X) --> topological_space(X::'a,Vt)) &  | 
|
2754  | 
(\<forall>Vt X Vs. finer(Vt::'a,Vs,X) --> topological_space(X::'a,Vs)) &  | 
|
2755  | 
(\<forall>X Vs Vt. finer(Vt::'a,Vs,X) --> subset_collections(Vs::'a,Vt)) &  | 
|
2756  | 
(\<forall>X Vs Vt. topological_space(X::'a,Vt) & topological_space(X::'a,Vs) & subset_collections(Vs::'a,Vt) --> finer(Vt::'a,Vs,X)) &  | 
|
2757  | 
(\<forall>Vf X. basis(X::'a,Vf) --> equal_sets(union_of_members(Vf),X)) &  | 
|
2758  | 
(\<forall>X Vf Y Vb1 Vb2. basis(X::'a,Vf) & element_of_set(Y::'a,X) & element_of_collection(Vb1::'a,Vf) & element_of_collection(Vb2::'a,Vf) & element_of_set(Y::'a,intersection_of_sets(Vb1::'a,Vb2)) --> element_of_set(Y::'a,f6(X::'a,Vf,Y,Vb1,Vb2))) &  | 
|
2759  | 
(\<forall>X Y Vb1 Vb2 Vf. basis(X::'a,Vf) & element_of_set(Y::'a,X) & element_of_collection(Vb1::'a,Vf) & element_of_collection(Vb2::'a,Vf) & element_of_set(Y::'a,intersection_of_sets(Vb1::'a,Vb2)) --> element_of_collection(f6(X::'a,Vf,Y,Vb1,Vb2),Vf)) &  | 
|
2760  | 
(\<forall>X Vf Y Vb1 Vb2. basis(X::'a,Vf) & element_of_set(Y::'a,X) & element_of_collection(Vb1::'a,Vf) & element_of_collection(Vb2::'a,Vf) & element_of_set(Y::'a,intersection_of_sets(Vb1::'a,Vb2)) --> subset_sets(f6(X::'a,Vf,Y,Vb1,Vb2),intersection_of_sets(Vb1::'a,Vb2))) &  | 
|
2761  | 
(\<forall>Vf X. equal_sets(union_of_members(Vf),X) --> basis(X::'a,Vf) | element_of_set(f7(X::'a,Vf),X)) &  | 
|
2762  | 
(\<forall>X Vf. equal_sets(union_of_members(Vf),X) --> basis(X::'a,Vf) | element_of_collection(f8(X::'a,Vf),Vf)) &  | 
|
2763  | 
(\<forall>X Vf. equal_sets(union_of_members(Vf),X) --> basis(X::'a,Vf) | element_of_collection(f9(X::'a,Vf),Vf)) &  | 
|
2764  | 
(\<forall>X Vf. equal_sets(union_of_members(Vf),X) --> basis(X::'a,Vf) | element_of_set(f7(X::'a,Vf),intersection_of_sets(f8(X::'a,Vf),f9(X::'a,Vf)))) &  | 
|
2765  | 
(\<forall>Uu9 X Vf. equal_sets(union_of_members(Vf),X) & element_of_set(f7(X::'a,Vf),Uu9) & element_of_collection(Uu9::'a,Vf) & subset_sets(Uu9::'a,intersection_of_sets(f8(X::'a,Vf),f9(X::'a,Vf))) --> basis(X::'a,Vf)) &  | 
|
2766  | 
(\<forall>Vf U X. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> element_of_set(X::'a,f10(Vf::'a,U,X))) &  | 
|
2767  | 
(\<forall>U X Vf. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> element_of_collection(f10(Vf::'a,U,X),Vf)) &  | 
|
2768  | 
(\<forall>Vf X U. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> subset_sets(f10(Vf::'a,U,X),U)) &  | 
|
2769  | 
(\<forall>Vf U. element_of_collection(U::'a,top_of_basis(Vf)) | element_of_set(f11(Vf::'a,U),U)) &  | 
|
2770  | 
(\<forall>Vf Uu11 U. element_of_set(f11(Vf::'a,U),Uu11) & element_of_collection(Uu11::'a,Vf) & subset_sets(Uu11::'a,U) --> element_of_collection(U::'a,top_of_basis(Vf))) &  | 
|
2771  | 
(\<forall>U Y X Vt. element_of_collection(U::'a,subspace_topology(X::'a,Vt,Y)) --> topological_space(X::'a,Vt)) &  | 
|
2772  | 
(\<forall>U Vt Y X. element_of_collection(U::'a,subspace_topology(X::'a,Vt,Y)) --> subset_sets(Y::'a,X)) &  | 
|
2773  | 
(\<forall>X Y U Vt. element_of_collection(U::'a,subspace_topology(X::'a,Vt,Y)) --> element_of_collection(f12(X::'a,Vt,Y,U),Vt)) &  | 
|
2774  | 
(\<forall>X Vt Y U. element_of_collection(U::'a,subspace_topology(X::'a,Vt,Y)) --> equal_sets(U::'a,intersection_of_sets(Y::'a,f12(X::'a,Vt,Y,U)))) &  | 
|
2775  | 
(\<forall>X Vt U Y Uu12. topological_space(X::'a,Vt) & subset_sets(Y::'a,X) & element_of_collection(Uu12::'a,Vt) & equal_sets(U::'a,intersection_of_sets(Y::'a,Uu12)) --> element_of_collection(U::'a,subspace_topology(X::'a,Vt,Y))) &  | 
|
2776  | 
(\<forall>U Y X Vt. element_of_set(U::'a,interior(Y::'a,X,Vt)) --> topological_space(X::'a,Vt)) &  | 
|
2777  | 
(\<forall>U Vt Y X. element_of_set(U::'a,interior(Y::'a,X,Vt)) --> subset_sets(Y::'a,X)) &  | 
|
2778  | 
(\<forall>Y X Vt U. element_of_set(U::'a,interior(Y::'a,X,Vt)) --> element_of_set(U::'a,f13(Y::'a,X,Vt,U))) &  | 
|
2779  | 
(\<forall>X Vt U Y. element_of_set(U::'a,interior(Y::'a,X,Vt)) --> subset_sets(f13(Y::'a,X,Vt,U),Y)) &  | 
|
2780  | 
(\<forall>Y U X Vt. element_of_set(U::'a,interior(Y::'a,X,Vt)) --> open(f13(Y::'a,X,Vt,U),X,Vt)) &  | 
|
2781  | 
(\<forall>U Y Uu13 X Vt. topological_space(X::'a,Vt) & subset_sets(Y::'a,X) & element_of_set(U::'a,Uu13) & subset_sets(Uu13::'a,Y) & open(Uu13::'a,X,Vt) --> element_of_set(U::'a,interior(Y::'a,X,Vt))) &  | 
|
2782  | 
(\<forall>U Y X Vt. element_of_set(U::'a,closure(Y::'a,X,Vt)) --> topological_space(X::'a,Vt)) &  | 
|
2783  | 
(\<forall>U Vt Y X. element_of_set(U::'a,closure(Y::'a,X,Vt)) --> subset_sets(Y::'a,X)) &  | 
|
2784  | 
(\<forall>Y X Vt U V. element_of_set(U::'a,closure(Y::'a,X,Vt)) & subset_sets(Y::'a,V) & closed(V::'a,X,Vt) --> element_of_set(U::'a,V)) &  | 
|
2785  | 
(\<forall>Y X Vt U. topological_space(X::'a,Vt) & subset_sets(Y::'a,X) --> element_of_set(U::'a,closure(Y::'a,X,Vt)) | subset_sets(Y::'a,f14(Y::'a,X,Vt,U))) &  | 
|
2786  | 
(\<forall>Y U X Vt. topological_space(X::'a,Vt) & subset_sets(Y::'a,X) --> element_of_set(U::'a,closure(Y::'a,X,Vt)) | closed(f14(Y::'a,X,Vt,U),X,Vt)) &  | 
|
2787  | 
(\<forall>Y X Vt U. topological_space(X::'a,Vt) & subset_sets(Y::'a,X) & element_of_set(U::'a,f14(Y::'a,X,Vt,U)) --> element_of_set(U::'a,closure(Y::'a,X,Vt))) &  | 
|
2788  | 
(\<forall>U Y X Vt. neighborhood(U::'a,Y,X,Vt) --> topological_space(X::'a,Vt)) &  | 
|
2789  | 
(\<forall>Y U X Vt. neighborhood(U::'a,Y,X,Vt) --> open(U::'a,X,Vt)) &  | 
|
2790  | 
(\<forall>X Vt Y U. neighborhood(U::'a,Y,X,Vt) --> element_of_set(Y::'a,U)) &  | 
|
2791  | 
(\<forall>X Vt Y U. topological_space(X::'a,Vt) & open(U::'a,X,Vt) & element_of_set(Y::'a,U) --> neighborhood(U::'a,Y,X,Vt)) &  | 
|
2792  | 
(\<forall>Z Y X Vt. limit_point(Z::'a,Y,X,Vt) --> topological_space(X::'a,Vt)) &  | 
|
2793  | 
(\<forall>Z Vt Y X. limit_point(Z::'a,Y,X,Vt) --> subset_sets(Y::'a,X)) &  | 
|
2794  | 
(\<forall>Z X Vt U Y. limit_point(Z::'a,Y,X,Vt) & neighborhood(U::'a,Z,X,Vt) --> element_of_set(f15(Z::'a,Y,X,Vt,U),intersection_of_sets(U::'a,Y))) &  | 
|
2795  | 
(\<forall>Y X Vt U Z. ~(limit_point(Z::'a,Y,X,Vt) & neighborhood(U::'a,Z,X,Vt) & eq_p(f15(Z::'a,Y,X,Vt,U),Z))) &  | 
|
2796  | 
(\<forall>Y Z X Vt. topological_space(X::'a,Vt) & subset_sets(Y::'a,X) --> limit_point(Z::'a,Y,X,Vt) | neighborhood(f16(Z::'a,Y,X,Vt),Z,X,Vt)) &  | 
|
2797  | 
(\<forall>X Vt Y Uu16 Z. topological_space(X::'a,Vt) & subset_sets(Y::'a,X) & element_of_set(Uu16::'a,intersection_of_sets(f16(Z::'a,Y,X,Vt),Y)) --> limit_point(Z::'a,Y,X,Vt) | eq_p(Uu16::'a,Z)) &  | 
|
2798  | 
(\<forall>U Y X Vt. element_of_set(U::'a,boundary(Y::'a,X,Vt)) --> topological_space(X::'a,Vt)) &  | 
|
2799  | 
(\<forall>U Y X Vt. element_of_set(U::'a,boundary(Y::'a,X,Vt)) --> element_of_set(U::'a,closure(Y::'a,X,Vt))) &  | 
|
2800  | 
(\<forall>U Y X Vt. element_of_set(U::'a,boundary(Y::'a,X,Vt)) --> element_of_set(U::'a,closure(relative_complement_sets(Y::'a,X),X,Vt))) &  | 
|
2801  | 
(\<forall>U Y X Vt. topological_space(X::'a,Vt) & element_of_set(U::'a,closure(Y::'a,X,Vt)) & element_of_set(U::'a,closure(relative_complement_sets(Y::'a,X),X,Vt)) --> element_of_set(U::'a,boundary(Y::'a,X,Vt))) &  | 
|
2802  | 
(\<forall>X Vt. hausdorff(X::'a,Vt) --> topological_space(X::'a,Vt)) &  | 
|
2803  | 
(\<forall>X_2 X_1 X Vt. hausdorff(X::'a,Vt) & element_of_set(X_1::'a,X) & element_of_set(X_2::'a,X) --> eq_p(X_1::'a,X_2) | neighborhood(f17(X::'a,Vt,X_1,X_2),X_1,X,Vt)) &  | 
|
2804  | 
(\<forall>X_1 X_2 X Vt. hausdorff(X::'a,Vt) & element_of_set(X_1::'a,X) & element_of_set(X_2::'a,X) --> eq_p(X_1::'a,X_2) | neighborhood(f18(X::'a,Vt,X_1,X_2),X_2,X,Vt)) &  | 
|
2805  | 
(\<forall>X Vt X_1 X_2. hausdorff(X::'a,Vt) & element_of_set(X_1::'a,X) & element_of_set(X_2::'a,X) --> eq_p(X_1::'a,X_2) | disjoint_s(f17(X::'a,Vt,X_1,X_2),f18(X::'a,Vt,X_1,X_2))) &  | 
|
2806  | 
(\<forall>Vt X. topological_space(X::'a,Vt) --> hausdorff(X::'a,Vt) | element_of_set(f19(X::'a,Vt),X)) &  | 
|
2807  | 
(\<forall>Vt X. topological_space(X::'a,Vt) --> hausdorff(X::'a,Vt) | element_of_set(f20(X::'a,Vt),X)) &  | 
|
2808  | 
(\<forall>X Vt. topological_space(X::'a,Vt) & eq_p(f19(X::'a,Vt),f20(X::'a,Vt)) --> hausdorff(X::'a,Vt)) &  | 
|
2809  | 
(\<forall>X Vt Uu19 Uu20. topological_space(X::'a,Vt) & neighborhood(Uu19::'a,f19(X::'a,Vt),X,Vt) & neighborhood(Uu20::'a,f20(X::'a,Vt),X,Vt) & disjoint_s(Uu19::'a,Uu20) --> hausdorff(X::'a,Vt)) &  | 
|
2810  | 
(\<forall>Va1 Va2 X Vt. separation(Va1::'a,Va2,X,Vt) --> topological_space(X::'a,Vt)) &  | 
|
2811  | 
(\<forall>Va2 X Vt Va1. ~(separation(Va1::'a,Va2,X,Vt) & equal_sets(Va1::'a,empty_set))) &  | 
|
2812  | 
(\<forall>Va1 X Vt Va2. ~(separation(Va1::'a,Va2,X,Vt) & equal_sets(Va2::'a,empty_set))) &  | 
|
2813  | 
(\<forall>Va2 X Va1 Vt. separation(Va1::'a,Va2,X,Vt) --> element_of_collection(Va1::'a,Vt)) &  | 
|
2814  | 
(\<forall>Va1 X Va2 Vt. separation(Va1::'a,Va2,X,Vt) --> element_of_collection(Va2::'a,Vt)) &  | 
|
2815  | 
(\<forall>Vt Va1 Va2 X. separation(Va1::'a,Va2,X,Vt) --> equal_sets(union_of_sets(Va1::'a,Va2),X)) &  | 
|
2816  | 
(\<forall>X Vt Va1 Va2. separation(Va1::'a,Va2,X,Vt) --> disjoint_s(Va1::'a,Va2)) &  | 
|
2817  | 
(\<forall>Vt X Va1 Va2. topological_space(X::'a,Vt) & element_of_collection(Va1::'a,Vt) & element_of_collection(Va2::'a,Vt) & equal_sets(union_of_sets(Va1::'a,Va2),X) & disjoint_s(Va1::'a,Va2) --> separation(Va1::'a,Va2,X,Vt) | equal_sets(Va1::'a,empty_set) | equal_sets(Va2::'a,empty_set)) &  | 
|
2818  | 
(\<forall>X Vt. connected_space(X::'a,Vt) --> topological_space(X::'a,Vt)) &  | 
|
2819  | 
(\<forall>Va1 Va2 X Vt. ~(connected_space(X::'a,Vt) & separation(Va1::'a,Va2,X,Vt))) &  | 
|
2820  | 
(\<forall>X Vt. topological_space(X::'a,Vt) --> connected_space(X::'a,Vt) | separation(f21(X::'a,Vt),f22(X::'a,Vt),X,Vt)) &  | 
|
2821  | 
(\<forall>Va X Vt. connected_set(Va::'a,X,Vt) --> topological_space(X::'a,Vt)) &  | 
|
2822  | 
(\<forall>Vt Va X. connected_set(Va::'a,X,Vt) --> subset_sets(Va::'a,X)) &  | 
|
2823  | 
(\<forall>X Vt Va. connected_set(Va::'a,X,Vt) --> connected_space(Va::'a,subspace_topology(X::'a,Vt,Va))) &  | 
|
2824  | 
(\<forall>X Vt Va. topological_space(X::'a,Vt) & subset_sets(Va::'a,X) & connected_space(Va::'a,subspace_topology(X::'a,Vt,Va)) --> connected_set(Va::'a,X,Vt)) &  | 
|
2825  | 
(\<forall>Vf X Vt. open_covering(Vf::'a,X,Vt) --> topological_space(X::'a,Vt)) &  | 
|
2826  | 
(\<forall>X Vf Vt. open_covering(Vf::'a,X,Vt) --> subset_collections(Vf::'a,Vt)) &  | 
|
2827  | 
(\<forall>Vt Vf X. open_covering(Vf::'a,X,Vt) --> equal_sets(union_of_members(Vf),X)) &  | 
|
2828  | 
(\<forall>Vt Vf X. topological_space(X::'a,Vt) & subset_collections(Vf::'a,Vt) & equal_sets(union_of_members(Vf),X) --> open_covering(Vf::'a,X,Vt)) &  | 
|
2829  | 
(\<forall>X Vt. compact_space(X::'a,Vt) --> topological_space(X::'a,Vt)) &  | 
|
2830  | 
(\<forall>X Vt Vf1. compact_space(X::'a,Vt) & open_covering(Vf1::'a,X,Vt) --> finite'(f23(X::'a,Vt,Vf1))) &  | 
|
2831  | 
(\<forall>X Vt Vf1. compact_space(X::'a,Vt) & open_covering(Vf1::'a,X,Vt) --> subset_collections(f23(X::'a,Vt,Vf1),Vf1)) &  | 
|
2832  | 
(\<forall>Vf1 X Vt. compact_space(X::'a,Vt) & open_covering(Vf1::'a,X,Vt) --> open_covering(f23(X::'a,Vt,Vf1),X,Vt)) &  | 
|
2833  | 
(\<forall>X Vt. topological_space(X::'a,Vt) --> compact_space(X::'a,Vt) | open_covering(f24(X::'a,Vt),X,Vt)) &  | 
|
2834  | 
(\<forall>Uu24 X Vt. topological_space(X::'a,Vt) & finite'(Uu24) & subset_collections(Uu24::'a,f24(X::'a,Vt)) & open_covering(Uu24::'a,X,Vt) --> compact_space(X::'a,Vt)) &  | 
|
2835  | 
(\<forall>Va X Vt. compact_set(Va::'a,X,Vt) --> topological_space(X::'a,Vt)) &  | 
|
2836  | 
(\<forall>Vt Va X. compact_set(Va::'a,X,Vt) --> subset_sets(Va::'a,X)) &  | 
|
2837  | 
(\<forall>X Vt Va. compact_set(Va::'a,X,Vt) --> compact_space(Va::'a,subspace_topology(X::'a,Vt,Va))) &  | 
|
2838  | 
(\<forall>X Vt Va. topological_space(X::'a,Vt) & subset_sets(Va::'a,X) & compact_space(Va::'a,subspace_topology(X::'a,Vt,Va)) --> compact_set(Va::'a,X,Vt)) &  | 
|
2839  | 
(basis(cx::'a,f)) &  | 
|
2840  | 
(\<forall>U. element_of_collection(U::'a,top_of_basis(f))) &  | 
|
2841  | 
(\<forall>V. element_of_collection(V::'a,top_of_basis(f))) &  | 
|
| 24127 | 2842  | 
(\<forall>U V. ~element_of_collection(intersection_of_sets(U::'a,V),top_of_basis(f))) --> False"  | 
2843  | 
by meson  | 
|
2844  | 
||
2845  | 
||
2846  | 
(*0 inferences so far. Searching to depth 0. 0.8 secs*)  | 
|
2847  | 
lemma TOP004_2:  | 
|
| 24128 | 2848  | 
"(\<forall>U Uu1 Vf. element_of_set(U::'a,Uu1) & element_of_collection(Uu1::'a,Vf) --> element_of_set(U::'a,union_of_members(Vf))) &  | 
2849  | 
(\<forall>Vf X. basis(X::'a,Vf) --> equal_sets(union_of_members(Vf),X)) &  | 
|
2850  | 
(\<forall>X Vf Y Vb1 Vb2. basis(X::'a,Vf) & element_of_set(Y::'a,X) & element_of_collection(Vb1::'a,Vf) & element_of_collection(Vb2::'a,Vf) & element_of_set(Y::'a,intersection_of_sets(Vb1::'a,Vb2)) --> element_of_set(Y::'a,f6(X::'a,Vf,Y,Vb1,Vb2))) &  | 
|
2851  | 
(\<forall>X Y Vb1 Vb2 Vf. basis(X::'a,Vf) & element_of_set(Y::'a,X) & element_of_collection(Vb1::'a,Vf) & element_of_collection(Vb2::'a,Vf) & element_of_set(Y::'a,intersection_of_sets(Vb1::'a,Vb2)) --> element_of_collection(f6(X::'a,Vf,Y,Vb1,Vb2),Vf)) &  | 
|
2852  | 
(\<forall>X Vf Y Vb1 Vb2. basis(X::'a,Vf) & element_of_set(Y::'a,X) & element_of_collection(Vb1::'a,Vf) & element_of_collection(Vb2::'a,Vf) & element_of_set(Y::'a,intersection_of_sets(Vb1::'a,Vb2)) --> subset_sets(f6(X::'a,Vf,Y,Vb1,Vb2),intersection_of_sets(Vb1::'a,Vb2))) &  | 
|
2853  | 
(\<forall>Vf U X. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> element_of_set(X::'a,f10(Vf::'a,U,X))) &  | 
|
2854  | 
(\<forall>U X Vf. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> element_of_collection(f10(Vf::'a,U,X),Vf)) &  | 
|
2855  | 
(\<forall>Vf X U. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> subset_sets(f10(Vf::'a,U,X),U)) &  | 
|
2856  | 
(\<forall>Vf U. element_of_collection(U::'a,top_of_basis(Vf)) | element_of_set(f11(Vf::'a,U),U)) &  | 
|
2857  | 
(\<forall>Vf Uu11 U. element_of_set(f11(Vf::'a,U),Uu11) & element_of_collection(Uu11::'a,Vf) & subset_sets(Uu11::'a,U) --> element_of_collection(U::'a,top_of_basis(Vf))) &  | 
|
2858  | 
(\<forall>Y X Z. subset_sets(X::'a,Y) & subset_sets(Y::'a,Z) --> subset_sets(X::'a,Z)) &  | 
|
2859  | 
(\<forall>Y Z X. element_of_set(Z::'a,intersection_of_sets(X::'a,Y)) --> element_of_set(Z::'a,X)) &  | 
|
2860  | 
(\<forall>X Z Y. element_of_set(Z::'a,intersection_of_sets(X::'a,Y)) --> element_of_set(Z::'a,Y)) &  | 
|
2861  | 
(\<forall>X Z Y. element_of_set(Z::'a,X) & element_of_set(Z::'a,Y) --> element_of_set(Z::'a,intersection_of_sets(X::'a,Y))) &  | 
|
2862  | 
(\<forall>X U Y V. subset_sets(X::'a,Y) & subset_sets(U::'a,V) --> subset_sets(intersection_of_sets(X::'a,U),intersection_of_sets(Y::'a,V))) &  | 
|
2863  | 
(\<forall>X Z Y. equal_sets(X::'a,Y) & element_of_set(Z::'a,X) --> element_of_set(Z::'a,Y)) &  | 
|
2864  | 
(\<forall>Y X. equal_sets(intersection_of_sets(X::'a,Y),intersection_of_sets(Y::'a,X))) &  | 
|
2865  | 
(basis(cx::'a,f)) &  | 
|
2866  | 
(\<forall>U. element_of_collection(U::'a,top_of_basis(f))) &  | 
|
2867  | 
(\<forall>V. element_of_collection(V::'a,top_of_basis(f))) &  | 
|
| 24127 | 2868  | 
(\<forall>U V. ~element_of_collection(intersection_of_sets(U::'a,V),top_of_basis(f))) --> False"  | 
2869  | 
by meson  | 
|
2870  | 
||
2871  | 
(*53777 inferences so far. Searching to depth 20. 68.7 secs*)  | 
|
2872  | 
lemma TOP005_2:  | 
|
| 24128 | 2873  | 
"(\<forall>Vf U. element_of_set(U::'a,union_of_members(Vf)) --> element_of_set(U::'a,f1(Vf::'a,U))) &  | 
2874  | 
(\<forall>U Vf. element_of_set(U::'a,union_of_members(Vf)) --> element_of_collection(f1(Vf::'a,U),Vf)) &  | 
|
2875  | 
(\<forall>Vf U X. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> element_of_set(X::'a,f10(Vf::'a,U,X))) &  | 
|
2876  | 
(\<forall>U X Vf. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> element_of_collection(f10(Vf::'a,U,X),Vf)) &  | 
|
2877  | 
(\<forall>Vf X U. element_of_collection(U::'a,top_of_basis(Vf)) & element_of_set(X::'a,U) --> subset_sets(f10(Vf::'a,U,X),U)) &  | 
|
2878  | 
(\<forall>Vf U. element_of_collection(U::'a,top_of_basis(Vf)) | element_of_set(f11(Vf::'a,U),U)) &  | 
|
2879  | 
(\<forall>Vf Uu11 U. element_of_set(f11(Vf::'a,U),Uu11) & element_of_collection(Uu11::'a,Vf) & subset_sets(Uu11::'a,U) --> element_of_collection(U::'a,top_of_basis(Vf))) &  | 
|
2880  | 
(\<forall>X U Y. element_of_set(U::'a,X) --> subset_sets(X::'a,Y) | element_of_set(U::'a,Y)) &  | 
|
2881  | 
(\<forall>Y X Z. subset_sets(X::'a,Y) & element_of_collection(Y::'a,Z) --> subset_sets(X::'a,union_of_members(Z))) &  | 
|
2882  | 
(\<forall>X U Y. subset_collections(X::'a,Y) & element_of_collection(U::'a,X) --> element_of_collection(U::'a,Y)) &  | 
|
2883  | 
(subset_collections(g::'a,top_of_basis(f))) &  | 
|
| 24127 | 2884  | 
(~element_of_collection(union_of_members(g),top_of_basis(f))) --> False"  | 
2885  | 
oops  | 
|
2886  | 
||
2887  | 
end  |