src/HOL/ex/mesontest2.ML
author paulson
Wed, 08 Oct 2003 15:57:41 +0200
changeset 14220 4dc132902672
parent 14183 466a2a69e7e8
child 15285 ce83b7e74a91
permissions -rw-r--r--
Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
     1
(*  Title:      HOL/ex/mesontest2
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
     2
    ID:         $Id$
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
     4
    Copyright   2000  University of Cambridge
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
     5
14220
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
     6
Test data for the MESON proof procedure
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
     7
   (Excludes the equality problems 51, 52, 56, 58)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
     8
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
     9
NOTE: most of the old file "mesontest.ML" has been converted to Isar and moved
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    10
to Classical.thy
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    11
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    12
Use the "mesonlog" shell script to process logs.
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    13
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    14
show_hyps := false;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    15
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    16
proofs := 0;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    17
by (rtac ccontr 1);
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    18
val [prem] = gethyps 1;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    19
val nnf = make_nnf prem;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    20
val xsko = skolemize nnf;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    21
by (cut_facts_tac [xsko] 1 THEN REPEAT (etac exE 1));
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    22
val [_,sko] = gethyps 1;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    23
val clauses = make_clauses [sko];       
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    24
val horns = make_horns clauses;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    25
val goes = gocls clauses;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    26
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    27
Goal "False";
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    28
by (resolve_tac goes 1);
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    29
proofs := 2;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    30
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    31
by (prolog_step_tac horns 1);
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    32
by (iter_deepen_prolog_tac horns);
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    33
by (depth_prolog_tac horns);
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    34
by (best_prolog_tac size_of_subgoals horns);
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    35
*)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    36
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    37
writeln"File HOL/ex/meson-test.";
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    38
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    39
context Main.thy;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    40
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    41
(*Deep unifications can be required, esp. during transformation to clauses*)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    42
val orig_trace_bound = !Unify.trace_bound
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    43
and orig_search_bound = !Unify.search_bound;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    44
Unify.trace_bound := 20;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    45
Unify.search_bound := 40;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    46
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    47
(**** Interactive examples ****)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    48
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    49
(*Generate nice names for Skolem functions*)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    50
Logic.auto_rename := true; Logic.set_rename_prefix "a";
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    51
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    52
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    53
writeln"Problem 25";
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    54
Goal "(\\<exists>x. P x) &  \
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    55
\     (\\<forall>x. L x --> ~ (M x & R x)) &  \
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    56
\     (\\<forall>x. P x --> (M x & L x)) &   \
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    57
\     ((\\<forall>x. P x --> Q x) | (\\<exists>x. P x & R x))  \
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    58
\   --> (\\<exists>x. Q x & P x)";
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    59
by (rtac ccontr 1);
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    60
val [prem25] = gethyps 1;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    61
val nnf25 = make_nnf prem25;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    62
val xsko25 = skolemize nnf25;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    63
by (cut_facts_tac [xsko25] 1 THEN REPEAT (etac exE 1));
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    64
val [_,sko25] = gethyps 1;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    65
val clauses25 = make_clauses [sko25];   (*7 clauses*)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    66
val horns25 = make_horns clauses25;     (*16 Horn clauses*)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    67
val go25::_ = gocls clauses25;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    68
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    69
Goal "False";
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    70
by (rtac go25 1);
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    71
by (depth_prolog_tac horns25);
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    72
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    73
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    74
writeln"Problem 26";
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    75
Goal "((\\<exists>x. p x) = (\\<exists>x. q x)) &     \
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    76
\     (\\<forall>x. \\<forall>y. p x & q y --> (r x = s y)) \
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    77
\ --> ((\\<forall>x. p x --> r x) = (\\<forall>x. q x --> s x))";
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    78
by (rtac ccontr 1);
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    79
val [prem26] = gethyps 1;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    80
val nnf26 = make_nnf prem26;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    81
val xsko26 = skolemize nnf26;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    82
by (cut_facts_tac [xsko26] 1 THEN REPEAT (etac exE 1));
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    83
val [_,sko26] = gethyps 1;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    84
val clauses26 = make_clauses [sko26];                   (*9 clauses*)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    85
val horns26 = make_horns clauses26;                     (*24 Horn clauses*)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    86
val go26::_ = gocls clauses26;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    87
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    88
Goal "False";
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    89
by (rtac go26 1);
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    90
by (depth_prolog_tac horns26);  (*1.4 secs*)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    91
(*Proof is of length 107!!*)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    92
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    93
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    94
writeln"Problem 43  NOW PROVED AUTOMATICALLY!!";  (*16 Horn clauses*)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    95
Goal "(\\<forall>x. \\<forall>y. q x y = (\\<forall>z. p z x = (p z y::bool)))  \
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    96
\     --> (\\<forall>x. (\\<forall>y. q x y = (q y x::bool)))";
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    97
by (rtac ccontr 1);
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    98
val [prem43] = gethyps 1;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
    99
val nnf43 = make_nnf prem43;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   100
val xsko43 = skolemize nnf43;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   101
by (cut_facts_tac [xsko43] 1 THEN REPEAT (etac exE 1));
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   102
val [_,sko43] = gethyps 1;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   103
val clauses43 = make_clauses [sko43];   (*6*)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   104
val horns43 = make_horns clauses43;     (*16*)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   105
val go43::_ = gocls clauses43;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   106
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   107
Goal "False";
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   108
by (rtac go43 1);
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   109
by (best_prolog_tac size_of_subgoals horns43);   (*1.6 secs*)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   110
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   111
(* 
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   112
#1  (q x xa ==> ~ q x xa) ==> q xa x
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   113
#2  (q xa x ==> ~ q xa x) ==> q x xa
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   114
#3  (~ q x xa ==> q x xa) ==> ~ q xa x
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   115
#4  (~ q xa x ==> q xa x) ==> ~ q x xa
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   116
#5  [| ~ q ?U ?V ==> q ?U ?V; ~ p ?W ?U ==> p ?W ?U |] ==> p ?W ?V
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   117
#6  [| ~ p ?W ?U ==> p ?W ?U; p ?W ?V ==> ~ p ?W ?V |] ==> ~ q ?U ?V
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   118
#7  [| p ?W ?V ==> ~ p ?W ?V; ~ q ?U ?V ==> q ?U ?V |] ==> ~ p ?W ?U
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   119
#8  [| ~ q ?U ?V ==> q ?U ?V; ~ p ?W ?V ==> p ?W ?V |] ==> p ?W ?U
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   120
#9  [| ~ p ?W ?V ==> p ?W ?V; p ?W ?U ==> ~ p ?W ?U |] ==> ~ q ?U ?V
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   121
#10 [| p ?W ?U ==> ~ p ?W ?U; ~ q ?U ?V ==> q ?U ?V |] ==> ~ p ?W ?V
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   122
#11 [| p (xb ?U ?V) ?U ==> ~ p (xb ?U ?V) ?U;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   123
       p (xb ?U ?V) ?V ==> ~ p (xb ?U ?V) ?V |] ==> q ?U ?V
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   124
#12 [| p (xb ?U ?V) ?V ==> ~ p (xb ?U ?V) ?V; q ?U ?V ==> ~ q ?U ?V |] ==>
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   125
    p (xb ?U ?V) ?U
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   126
#13 [| q ?U ?V ==> ~ q ?U ?V; p (xb ?U ?V) ?U ==> ~ p (xb ?U ?V) ?U |] ==>
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   127
    p (xb ?U ?V) ?V
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   128
#14 [| ~ p (xb ?U ?V) ?U ==> p (xb ?U ?V) ?U;
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   129
       ~ p (xb ?U ?V) ?V ==> p (xb ?U ?V) ?V |] ==> q ?U ?V
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   130
#15 [| ~ p (xb ?U ?V) ?V ==> p (xb ?U ?V) ?V; q ?U ?V ==> ~ q ?U ?V |] ==>
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   131
    ~ p (xb ?U ?V) ?U
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   132
#16 [| q ?U ?V ==> ~ q ?U ?V; ~ p (xb ?U ?V) ?U ==> p (xb ?U ?V) ?U |] ==>
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   133
    ~ p (xb ?U ?V) ?V
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   134
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   135
And here is the proof! (Unkn is the start state after use of goal clause)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   136
[Unkn, Res ([Thm "#14"], false, 1), Res ([Thm "#5"], false, 1),
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   137
   Res ([Thm "#1"], false, 1), Asm 1, Res ([Thm "#13"], false, 1), Asm 2,
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   138
   Asm 1, Res ([Thm "#13"], false, 1), Asm 1, Res ([Thm "#10"], false, 1),
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   139
   Res ([Thm "#16"], false, 1), Asm 2, Asm 1, Res ([Thm "#1"], false, 1),
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   140
   Asm 1, Res ([Thm "#14"], false, 1), Res ([Thm "#5"], false, 1),
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   141
   Res ([Thm "#2"], false, 1), Asm 1, Res ([Thm "#13"], false, 1), Asm 2,
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   142
   Asm 1, Res ([Thm "#8"], false, 1), Res ([Thm "#2"], false, 1), Asm 1,
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   143
   Res ([Thm "#12"], false, 1), Asm 2, Asm 1] : lderiv list
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   144
*)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   145
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   146
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   147
(*Restore variable name preservation*)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   148
Logic.auto_rename := false; 
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   149
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   150
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   151
(** Charles Morgan's problems **)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   152
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   153
val axa = "\\<forall>x y.   T(i x(i y x))";
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   154
val axb = "\\<forall>x y z. T(i(i x(i y z))(i(i x y)(i x z)))";
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   155
val axc = "\\<forall>x y.   T(i(i(n x)(n y))(i y x))";
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   156
val axd = "\\<forall>x y.   T(i x y) & T x --> T y";
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   157
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   158
fun axjoin ([],   q) = q
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   159
  | axjoin (p::ps, q) = "(" ^ p ^ ") --> (" ^ axjoin(ps,q) ^ ")";
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   160
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   161
Goal (axjoin([axa,axb,axd], "\\<forall>x. T(i x x)"));
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   162
by (meson_tac 1);  
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   163
result();
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   164
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   165
writeln"Problem 66";  
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   166
Goal (axjoin([axa,axb,axc,axd], "\\<forall>x. T(i x(n(n x)))"));
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   167
by (meson_tac 1);  
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   168
result();
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   169
(*SLOW: 37s on a 1.8MHz machine
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   170
     208346 inferences so far.  Searching to depth 23 *)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   171
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   172
writeln"Problem 67";  
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   173
Goal (axjoin([axa,axb,axc,axd], "\\<forall>x. T(i(n(n x)) x)"));
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   174
by (meson_tac 1);  
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   175
result();
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   176
(*10s on a 1.8MHz machine
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   177
  51061 inferences so far.  Searching to depth 21 *)
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   178
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   179
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   180
(*MORE and MUCH HARDER test data for the MESON proof procedure
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   181
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   182
Courtesy John Harrison 
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
   183
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
   184
WARNING: there are many potential conflicts between variables used below
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
   185
and constants declared in HOL!
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   186
*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   187
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   188
(*All but the fastest are ignored to reduce build time*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   189
val even_hard_ones = false;
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   190
10440
2074e62da354 proper theory context for mesontest2;
wenzelm
parents: 10212
diff changeset
   191
fun prove (s,tac) = prove_goal (the_context ()) s (fn [] => [tac]);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   192
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   193
fun prove_hard arg = if even_hard_ones then prove arg else TrueI;
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   194
9000
c20d58286a51 cleaned up;
wenzelm
parents: 8557
diff changeset
   195
set timing;
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   196
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
   197
context Main.thy;
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
   198
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   199
(* ========================================================================= *)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   200
(* 100 problems selected from the TPTP library                               *)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   201
(* ========================================================================= *)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   202
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   203
(*
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   204
 * Original timings for John Harrison's MESON_TAC.
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   205
 * Timings below on a 600MHz Pentium III (perch)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   206
 * 
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   207
 * A few variable names have been primed to avoid clashing with constants.
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   208
 *
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   209
 * Changed numeric constants e.g. 0, 1, 2... to num0, num1, num2...
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   210
 *
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   211
 * Here's a list giving typical CPU times, as well as common names and
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   212
 * literature references.
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   213
 *
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   214
 * BOO003-1     34.6    B2 part 1 [McCharen, et al., 1976]; Lemma proved [Overbeek, et al., 1976]; prob2_part1.ver1.in [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   215
 * BOO004-1     36.7    B2 part 2 [McCharen, et al., 1976]; Lemma proved [Overbeek, et al., 1976]; prob2_part2.ver1 [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   216
 * 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]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   217
 * 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]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   218
 * BOO011-1     19.0    B7 [McCharen, et al., 1976]; prob7.ver1 [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   219
 * CAT001-3     45.2    C1 [McCharen, et al., 1976]; p1.ver3.in [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   220
 * CAT003-3     10.5    C3 [McCharen, et al., 1976]; p3.ver3.in [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   221
 * CAT005-1    480.1    C5 [McCharen, et al., 1976]; p5.ver1.in [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   222
 * CAT007-1     11.9    C7 [McCharen, et al., 1976]; p7.ver1.in [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   223
 * CAT018-1     81.3    p18.ver1.in [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   224
 * COL001-2     16.0    C1 [Wos & McCune, 1988]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   225
 * COL023-1      5.1    [McCune & Wos, 1988]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   226
 * COL032-1     15.8    [McCune & Wos, 1988]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   227
 * COL052-2     13.2    bird4.ver2.in [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   228
 * COL075-2    116.9    [Jech, 1994]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   229
 * COM001-1      1.7    shortburst [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   230
 * COM002-1      4.4    burstall [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   231
 * COM002-2      7.4
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   232
 * COM003-2     22.1    [Brushi, 1991]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   233
 * COM004-1     45.1
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   234
 * GEO003-1     71.7    T3 [McCharen, et al., 1976]; t3.ver1.in [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   235
 * GEO017-2     78.8    D4.1 [Quaife, 1989]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   236
 * GEO027-3    181.5    D10.1 [Quaife, 1989]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   237
 * GEO058-2    104.0    R4 [Quaife, 1989]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   238
 * GEO079-1      2.4    GEOMETRY THEOREM [Slagle, 1967]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   239
 * 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]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   240
 * GRP008-1     50.4    Problem 4 [Wos, 1965]; wos4 [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   241
 * GRP013-1     40.2    Problem 11 [Wos, 1965]; wos11 [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   242
 * GRP037-3     43.8    Problem 17 [Wos, 1965]; wos17 [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   243
 * GRP031-2      3.2    ls23 [Lawrence & Starkey, 1974]; ls23 [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   244
 * GRP034-4      2.5    ls26 [Lawrence & Starkey, 1974]; ls26 [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   245
 * GRP047-2     11.7    [Veroff, 1992]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   246
 * GRP130-1    170.6    Bennett QG8 [TPTP]; QG8 [Slaney, 1993]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   247
 * GRP156-1     48.7    ax_mono1c [Schulz, 1995]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   248
 * GRP168-1    159.1    p01a [Schulz, 1995]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   249
 * HEN003-3     39.9    HP3 [McCharen, et al., 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   250
 * HEN007-2    125.7    H7 [McCharen, et al., 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   251
 * HEN008-4     62.0    H8 [McCharen, et al., 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   252
 * HEN009-5    136.3    H9 [McCharen, et al., 1976]; hp9.ver3.in [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   253
 * HEN012-3     48.5    new.ver2.in [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   254
 * LCL010-1    370.9    EC-73 [McCune & Wos, 1992]; ec_yq.in [OTTER]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   255
 * LCL077-2     51.6    morgan.two.ver1.in [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   256
 * LCL082-1     14.6    IC-1.1 [Wos, et al., 1990]; IC-65 [McCune & Wos, 1992]; ls2 [SETHEO]; S1 [Pfenning, 1988]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   257
 * 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]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   258
 * LCL143-1     10.9    Lattice structure theorem 2 [Bonacina, 1991]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   259
 * LCL182-1    271.6    Problem 2.16 [Whitehead & Russell, 1927]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   260
 * LCL200-1     12.0    Problem 2.46 [Whitehead & Russell, 1927]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   261
 * LCL215-1    214.4    Problem 2.62 [Whitehead & Russell, 1927]; Problem 2.63 [Whitehead & Russell, 1927]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   262
 * LCL230-2      0.2    Pelletier 5 [Pelletier, 1986]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   263
 * LDA003-1     68.5    Problem 3 [Jech, 1993]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   264
 * MSC002-1      9.2    DBABHP [Michie, et al., 1972]; DBABHP [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   265
 * MSC003-1      3.2    HASPARTS-T1 [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   266
 * MSC004-1      9.3    HASPARTS-T2 [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   267
 * MSC005-1      1.8    Problem 5.1 [Plaisted, 1982]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   268
 * MSC006-1     39.0    nonob.lop [SETHEO]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   269
 * NUM001-1     14.0    Chang-Lee-10a [Chang, 1970]; ls28 [Lawrence & Starkey, 1974]; ls28 [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   270
 * NUM021-1     52.3    ls65 [Lawrence & Starkey, 1974]; ls65 [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   271
 * NUM024-1     64.6    ls75 [Lawrence & Starkey, 1974]; ls75 [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   272
 * NUM180-1    621.2    LIM2.1 [Quaife]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   273
 * NUM228-1    575.9    TRECDEF4 cor. [Quaife]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   274
 * PLA002-1     37.4    Problem 5.7 [Plaisted, 1982]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   275
 * PLA006-1      7.2    [Segre & Elkan, 1994]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   276
 * PLA017-1    484.8    [Segre & Elkan, 1994]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   277
 * PLA022-1     19.1    [Segre & Elkan, 1994]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   278
 * PLA022-2     19.7    [Segre & Elkan, 1994]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   279
 * PRV001-1     10.3    PV1 [McCharen, et al., 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   280
 * PRV003-1      3.9    E2 [McCharen, et al., 1976]; v2.lop [SETHEO]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   281
 * PRV005-1      4.3    E4 [McCharen, et al., 1976]; v4.lop [SETHEO]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   282
 * PRV006-1      6.0    E5 [McCharen, et al., 1976]; v5.lop [SETHEO]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   283
 * PRV009-1      2.2    Hoares FIND [Bledsoe, 1977]; Problem 5.5 [Plaisted, 1982]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   284
 * PUZ012-1      3.5    Boxes-of-fruit [Wos, 1988]; Boxes-of-fruit [Wos, et al., 1992]; boxes.ver1.in [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   285
 * PUZ020-1     56.6    knightknave.in [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   286
 * PUZ025-1     58.4    Problem 35 [Smullyan, 1978]; tandl35.ver1.in [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   287
 * PUZ029-1      5.1    pigs.ver1.in [ANL]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   288
 * RNG001-3     82.4    EX6-T? [Wilson & Minker, 1976]; ex6.lop [SETHEO]; Example 6a [Fleisig, et al., 1974]; FEX6T1 [SPRFN]; FEX6T2 [SPRFN]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   289
 * RNG001-5    399.8    Problem 21 [Wos, 1965]; wos21 [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   290
 * RNG011-5      8.4    CADE-11 Competition Eq-10 [Overbeek, 1990]; PROBLEM 10 [Zhang, 1993]; THEOREM EQ-10 [Lusk & McCune, 1993]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   291
 * RNG023-6      9.1    [Stevens, 1987]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   292
 * RNG028-2      9.3    PROOF III [Anantharaman & Hsiang, 1990]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   293
 * RNG038-2     16.2    Problem 27 [Wos, 1965]; wos27 [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   294
 * RNG040-2    180.5    Problem 29 [Wos, 1965]; wos29 [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   295
 * RNG041-1     35.8    Problem 30 [Wos, 1965]; wos30 [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   296
 * ROB010-1    205.0    Lemma 3.3 [Winker, 1990]; RA2 [Lusk & Wos, 1992]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   297
 * ROB013-1     23.6    Lemma 3.5 [Winker, 1990]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   298
 * ROB016-1     15.2    Corollary 3.7 [Winker, 1990]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   299
 * ROB021-1    230.4    [McCune, 1992]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   300
 * SET005-1    192.2    ls108 [Lawrence & Starkey, 1974]; ls108 [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   301
 * SET009-1     10.5    ls116 [Lawrence & Starkey, 1974]; ls116 [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   302
 * SET025-4    694.7    Lemma 10 [Boyer, et al, 1986]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   303
 * SET046-5      2.3    p42.in [ANL]; Pelletier 42 [Pelletier, 1986]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   304
 * SET047-5      3.7    p43.in [ANL]; Pelletier 43 [Pelletier, 1986]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   305
 * SYN034-1      2.8    QW [Michie, et al., 1972]; QW [Wilson & Minker, 1976]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   306
 * SYN071-1      1.9    Pelletier 48 [Pelletier, 1986]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   307
 * SYN349-1     61.7    Ch17N5 [Tammet, 1994]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   308
 * SYN352-1      5.5    Ch18N4 [Tammet, 1994]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   309
 * TOP001-2     61.1    Lemma 1a [Wick & McCune, 1989]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   310
 * TOP002-2      0.4    Lemma 1b [Wick & McCune, 1989]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   311
 * TOP004-1    181.6    Lemma 1d [Wick & McCune, 1989]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   312
 * TOP004-2      9.0    Lemma 1d [Wick & McCune, 1989]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   313
 * TOP005-2    139.8    Lemma 1e [Wick & McCune, 1989]
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   314
 *)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   315
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   316
(*51194 inferences so far.  Searching to depth 13.  232.9 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   317
val BOO003_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   318
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   319
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   320
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   321
\  (\\<forall>X Y. sum(X::'a,Y,add(X::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   322
\  (\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   323
\  (\\<forall>Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   324
\  (\\<forall>Y X Z. product(X::'a,Y,Z) --> product(Y::'a,X,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   325
\  (\\<forall>X. sum(additive_identity::'a,X,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   326
\  (\\<forall>X. sum(X::'a,additive_identity,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   327
\  (\\<forall>X. product(multiplicative_identity::'a,X,X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   328
\  (\\<forall>X. product(X::'a,multiplicative_identity,X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   329
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   330
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   331
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   332
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   333
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   334
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   335
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   336
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   337
\  (\\<forall>X. sum(INVERSE(X),X,multiplicative_identity)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   338
\  (\\<forall>X. sum(X::'a,INVERSE(X),multiplicative_identity)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   339
\  (\\<forall>X. product(INVERSE(X),X,additive_identity)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   340
\  (\\<forall>X. product(X::'a,INVERSE(X),additive_identity)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   341
\  (\\<forall>X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   342
\  (\\<forall>X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   343
\  (\\<forall>X Y W Z. equal(X::'a,Y) & sum(X::'a,W,Z) --> sum(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   344
\  (\\<forall>X W Y Z. equal(X::'a,Y) & sum(W::'a,X,Z) --> sum(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   345
\  (\\<forall>X W Z Y. equal(X::'a,Y) & sum(W::'a,Z,X) --> sum(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   346
\  (\\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   347
\  (\\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   348
\  (\\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   349
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(add(X::'a,W),add(Y::'a,W))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   350
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(add(W::'a,X),add(W::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   351
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   352
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   353
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(INVERSE(X),INVERSE(Y))) &       \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   354
\  (~product(x::'a,x,x)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   355
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   356
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   357
(*51194 inferences so far.  Searching to depth 13. 204.6 secs
14220
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
   358
  Strange! The previous problem also has 51194 inferences at depth 13.  They
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   359
   must be very similar!*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   360
val BOO004_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   361
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   362
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   363
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   364
\  (\\<forall>X Y. sum(X::'a,Y,add(X::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   365
\  (\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   366
\  (\\<forall>Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   367
\  (\\<forall>Y X Z. product(X::'a,Y,Z) --> product(Y::'a,X,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   368
\  (\\<forall>X. sum(additive_identity::'a,X,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   369
\  (\\<forall>X. sum(X::'a,additive_identity,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   370
\  (\\<forall>X. product(multiplicative_identity::'a,X,X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   371
\  (\\<forall>X. product(X::'a,multiplicative_identity,X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   372
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   373
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   374
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   375
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   376
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   377
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   378
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   379
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   380
\  (\\<forall>X. sum(INVERSE(X),X,multiplicative_identity)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   381
\  (\\<forall>X. sum(X::'a,INVERSE(X),multiplicative_identity)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   382
\  (\\<forall>X. product(INVERSE(X),X,additive_identity)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   383
\  (\\<forall>X. product(X::'a,INVERSE(X),additive_identity)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   384
\  (\\<forall>X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   385
\  (\\<forall>X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   386
\  (\\<forall>X Y W Z. equal(X::'a,Y) & sum(X::'a,W,Z) --> sum(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   387
\  (\\<forall>X W Y Z. equal(X::'a,Y) & sum(W::'a,X,Z) --> sum(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   388
\  (\\<forall>X W Z Y. equal(X::'a,Y) & sum(W::'a,Z,X) --> sum(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   389
\  (\\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   390
\  (\\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   391
\  (\\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   392
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(add(X::'a,W),add(Y::'a,W))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   393
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(add(W::'a,X),add(W::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   394
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   395
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   396
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(INVERSE(X),INVERSE(Y))) &       \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   397
\  (~sum(x::'a,x,x)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   398
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   399
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   400
(*74799 inferences so far.  Searching to depth 13.  290.0 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   401
val BOO005_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   402
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   403
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   404
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   405
\  (\\<forall>X Y. sum(X::'a,Y,add(X::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   406
\  (\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   407
\  (\\<forall>Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   408
\  (\\<forall>Y X Z. product(X::'a,Y,Z) --> product(Y::'a,X,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   409
\  (\\<forall>X. sum(additive_identity::'a,X,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   410
\  (\\<forall>X. sum(X::'a,additive_identity,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   411
\  (\\<forall>X. product(multiplicative_identity::'a,X,X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   412
\  (\\<forall>X. product(X::'a,multiplicative_identity,X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   413
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   414
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   415
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   416
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   417
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   418
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   419
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   420
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   421
\  (\\<forall>X. sum(INVERSE(X),X,multiplicative_identity)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   422
\  (\\<forall>X. sum(X::'a,INVERSE(X),multiplicative_identity)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   423
\  (\\<forall>X. product(INVERSE(X),X,additive_identity)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   424
\  (\\<forall>X. product(X::'a,INVERSE(X),additive_identity)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   425
\  (\\<forall>X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   426
\  (\\<forall>X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   427
\  (\\<forall>X Y W Z. equal(X::'a,Y) & sum(X::'a,W,Z) --> sum(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   428
\  (\\<forall>X W Y Z. equal(X::'a,Y) & sum(W::'a,X,Z) --> sum(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   429
\  (\\<forall>X W Z Y. equal(X::'a,Y) & sum(W::'a,Z,X) --> sum(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   430
\  (\\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   431
\  (\\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   432
\  (\\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   433
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(add(X::'a,W),add(Y::'a,W))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   434
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(add(W::'a,X),add(W::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   435
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   436
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   437
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(INVERSE(X),INVERSE(Y))) &       \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   438
\  (~sum(x::'a,multiplicative_identity,multiplicative_identity)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   439
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   440
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   441
(*74799 inferences so far.  Searching to depth 13.  314.6 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   442
val BOO006_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   443
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   444
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   445
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   446
\  (\\<forall>X Y. sum(X::'a,Y,add(X::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   447
\  (\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   448
\  (\\<forall>Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   449
\  (\\<forall>Y X Z. product(X::'a,Y,Z) --> product(Y::'a,X,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   450
\  (\\<forall>X. sum(additive_identity::'a,X,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   451
\  (\\<forall>X. sum(X::'a,additive_identity,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   452
\  (\\<forall>X. product(multiplicative_identity::'a,X,X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   453
\  (\\<forall>X. product(X::'a,multiplicative_identity,X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   454
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   455
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   456
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   457
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   458
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   459
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   460
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   461
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   462
\  (\\<forall>X. sum(INVERSE(X),X,multiplicative_identity)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   463
\  (\\<forall>X. sum(X::'a,INVERSE(X),multiplicative_identity)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   464
\  (\\<forall>X. product(INVERSE(X),X,additive_identity)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   465
\  (\\<forall>X. product(X::'a,INVERSE(X),additive_identity)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   466
\  (\\<forall>X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   467
\  (\\<forall>X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   468
\  (\\<forall>X Y W Z. equal(X::'a,Y) & sum(X::'a,W,Z) --> sum(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   469
\  (\\<forall>X W Y Z. equal(X::'a,Y) & sum(W::'a,X,Z) --> sum(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   470
\  (\\<forall>X W Z Y. equal(X::'a,Y) & sum(W::'a,Z,X) --> sum(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   471
\  (\\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   472
\  (\\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   473
\  (\\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   474
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(add(X::'a,W),add(Y::'a,W))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   475
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(add(W::'a,X),add(W::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   476
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   477
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   478
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(INVERSE(X),INVERSE(Y))) &       \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   479
\  (~product(x::'a,additive_identity,additive_identity)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   480
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   481
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   482
(*5 inferences so far.  Searching to depth 5.  1.3 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   483
val BOO011_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   484
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   485
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   486
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   487
\  (\\<forall>X Y. sum(X::'a,Y,add(X::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   488
\  (\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   489
\  (\\<forall>Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   490
\  (\\<forall>Y X Z. product(X::'a,Y,Z) --> product(Y::'a,X,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   491
\  (\\<forall>X. sum(additive_identity::'a,X,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   492
\  (\\<forall>X. sum(X::'a,additive_identity,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   493
\  (\\<forall>X. product(multiplicative_identity::'a,X,X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   494
\  (\\<forall>X. product(X::'a,multiplicative_identity,X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   495
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   496
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   497
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   498
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   499
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   500
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   501
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   502
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   503
\  (\\<forall>X. sum(INVERSE(X),X,multiplicative_identity)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   504
\  (\\<forall>X. sum(X::'a,INVERSE(X),multiplicative_identity)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   505
\  (\\<forall>X. product(INVERSE(X),X,additive_identity)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   506
\  (\\<forall>X. product(X::'a,INVERSE(X),additive_identity)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   507
\  (\\<forall>X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   508
\  (\\<forall>X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   509
\  (\\<forall>X Y W Z. equal(X::'a,Y) & sum(X::'a,W,Z) --> sum(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   510
\  (\\<forall>X W Y Z. equal(X::'a,Y) & sum(W::'a,X,Z) --> sum(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   511
\  (\\<forall>X W Z Y. equal(X::'a,Y) & sum(W::'a,Z,X) --> sum(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   512
\  (\\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   513
\  (\\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   514
\  (\\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   515
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(add(X::'a,W),add(Y::'a,W))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   516
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(add(W::'a,X),add(W::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   517
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   518
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   519
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(INVERSE(X),INVERSE(Y))) &       \
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
   520
\  (~equal(INVERSE(additive_identity),multiplicative_identity)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   521
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   522
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   523
(*4007 inferences so far.  Searching to depth 9.  13 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   524
val CAT001_3 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   525
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   526
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   527
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   528
\  (\\<forall>Y X. equivalent(X::'a,Y) --> there_exists(X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   529
\  (\\<forall>X Y. equivalent(X::'a,Y) --> equal(X::'a,Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   530
\  (\\<forall>X Y. there_exists(X) & equal(X::'a,Y) --> equivalent(X::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   531
\  (\\<forall>X. there_exists(domain(X)) --> there_exists(X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   532
\  (\\<forall>X. there_exists(codomain(X)) --> there_exists(X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   533
\  (\\<forall>Y X. there_exists(compos(X::'a,Y)) --> there_exists(domain(X))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   534
\  (\\<forall>X Y. there_exists(compos(X::'a,Y)) --> equal(domain(X),codomain(Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   535
\  (\\<forall>X Y. there_exists(domain(X)) & equal(domain(X),codomain(Y)) --> there_exists(compos(X::'a,Y))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   536
\  (\\<forall>X Y Z. equal(compos(X::'a,compos(Y::'a,Z)),compos(compos(X::'a,Y),Z))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   537
\  (\\<forall>X. equal(compos(X::'a,domain(X)),X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   538
\  (\\<forall>X. equal(compos(codomain(X),X),X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   539
\  (\\<forall>X Y. equivalent(X::'a,Y) --> there_exists(Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   540
\  (\\<forall>X Y. there_exists(X) & there_exists(Y) & equal(X::'a,Y) --> equivalent(X::'a,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   541
\  (\\<forall>Y X. there_exists(compos(X::'a,Y)) --> there_exists(codomain(X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   542
\  (\\<forall>X Y. there_exists(f1(X::'a,Y)) | equal(X::'a,Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   543
\  (\\<forall>X Y. equal(X::'a,f1(X::'a,Y)) | equal(Y::'a,f1(X::'a,Y)) | equal(X::'a,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   544
\  (\\<forall>X Y. equal(X::'a,f1(X::'a,Y)) & equal(Y::'a,f1(X::'a,Y)) --> equal(X::'a,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   545
\  (\\<forall>X Y. equal(X::'a,Y) & there_exists(X) --> there_exists(Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   546
\  (\\<forall>X Y Z. equal(X::'a,Y) & equivalent(X::'a,Z) --> equivalent(Y::'a,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   547
\  (\\<forall>X Z Y. equal(X::'a,Y) & equivalent(Z::'a,X) --> equivalent(Z::'a,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   548
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(domain(X),domain(Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   549
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(codomain(X),codomain(Y))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   550
\  (\\<forall>X Y Z. equal(X::'a,Y) --> equal(compos(X::'a,Z),compos(Y::'a,Z))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   551
\  (\\<forall>X Z Y. equal(X::'a,Y) --> equal(compos(Z::'a,X),compos(Z::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   552
\  (\\<forall>A B C. equal(A::'a,B) --> equal(f1(A::'a,C),f1(B::'a,C))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   553
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(f1(F'::'a,D),f1(F'::'a,E))) &        \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   554
\  (there_exists(compos(a::'a,b))) &       \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   555
\  (\\<forall>Y X Z. equal(compos(compos(a::'a,b),X),Y) & equal(compos(compos(a::'a,b),Z),Y) --> equal(X::'a,Z)) &      \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   556
\  (there_exists(compos(b::'a,h))) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   557
\  (equal(compos(b::'a,h),compos(b::'a,g))) & \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   558
\  (~equal(h::'a,g)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   559
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   560
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   561
(*245 inferences so far.  Searching to depth 7.  1.0 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   562
val CAT003_3 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   563
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   564
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   565
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   566
\  (\\<forall>Y X. equivalent(X::'a,Y) --> there_exists(X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   567
\  (\\<forall>X Y. equivalent(X::'a,Y) --> equal(X::'a,Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   568
\  (\\<forall>X Y. there_exists(X) & equal(X::'a,Y) --> equivalent(X::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   569
\  (\\<forall>X. there_exists(domain(X)) --> there_exists(X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   570
\  (\\<forall>X. there_exists(codomain(X)) --> there_exists(X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   571
\  (\\<forall>Y X. there_exists(compos(X::'a,Y)) --> there_exists(domain(X))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   572
\  (\\<forall>X Y. there_exists(compos(X::'a,Y)) --> equal(domain(X),codomain(Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   573
\  (\\<forall>X Y. there_exists(domain(X)) & equal(domain(X),codomain(Y)) --> there_exists(compos(X::'a,Y))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   574
\  (\\<forall>X Y Z. equal(compos(X::'a,compos(Y::'a,Z)),compos(compos(X::'a,Y),Z))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   575
\  (\\<forall>X. equal(compos(X::'a,domain(X)),X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   576
\  (\\<forall>X. equal(compos(codomain(X),X),X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   577
\  (\\<forall>X Y. equivalent(X::'a,Y) --> there_exists(Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   578
\  (\\<forall>X Y. there_exists(X) & there_exists(Y) & equal(X::'a,Y) --> equivalent(X::'a,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   579
\  (\\<forall>Y X. there_exists(compos(X::'a,Y)) --> there_exists(codomain(X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   580
\  (\\<forall>X Y. there_exists(f1(X::'a,Y)) | equal(X::'a,Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   581
\  (\\<forall>X Y. equal(X::'a,f1(X::'a,Y)) | equal(Y::'a,f1(X::'a,Y)) | equal(X::'a,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   582
\  (\\<forall>X Y. equal(X::'a,f1(X::'a,Y)) & equal(Y::'a,f1(X::'a,Y)) --> equal(X::'a,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   583
\  (\\<forall>X Y. equal(X::'a,Y) & there_exists(X) --> there_exists(Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   584
\  (\\<forall>X Y Z. equal(X::'a,Y) & equivalent(X::'a,Z) --> equivalent(Y::'a,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   585
\  (\\<forall>X Z Y. equal(X::'a,Y) & equivalent(Z::'a,X) --> equivalent(Z::'a,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   586
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(domain(X),domain(Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   587
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(codomain(X),codomain(Y))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   588
\  (\\<forall>X Y Z. equal(X::'a,Y) --> equal(compos(X::'a,Z),compos(Y::'a,Z))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   589
\  (\\<forall>X Z Y. equal(X::'a,Y) --> equal(compos(Z::'a,X),compos(Z::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   590
\  (\\<forall>A B C. equal(A::'a,B) --> equal(f1(A::'a,C),f1(B::'a,C))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   591
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(f1(F'::'a,D),f1(F'::'a,E))) &        \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   592
\  (there_exists(compos(a::'a,b))) &       \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   593
\  (\\<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)) &      \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   594
\  (there_exists(h)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   595
\  (equal(compos(h::'a,a),compos(g::'a,a))) & \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   596
\  (~equal(g::'a,h)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   597
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   598
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   599
(*54288 inferences so far.  Searching to depth 14.  118.0 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   600
val CAT005_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   601
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   602
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   603
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   604
\  (\\<forall>X Y. defined(X::'a,Y) --> product(X::'a,Y,compos(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   605
\  (\\<forall>Z X Y. product(X::'a,Y,Z) --> defined(X::'a,Y)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   606
\  (\\<forall>X Xy Y Z. product(X::'a,Y,Xy) & defined(Xy::'a,Z) --> defined(Y::'a,Z)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   607
\  (\\<forall>Y Xy Z X Yz. product(X::'a,Y,Xy) & product(Y::'a,Z,Yz) & defined(Xy::'a,Z) --> defined(X::'a,Yz)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   608
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   609
\  (\\<forall>Z Yz X Y. product(Y::'a,Z,Yz) & defined(X::'a,Yz) --> defined(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   610
\  (\\<forall>Y X Yz Xy Z. product(Y::'a,Z,Yz) & product(X::'a,Y,Xy) & defined(X::'a,Yz) --> defined(Xy::'a,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   611
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   612
\  (\\<forall>Y X Z. defined(X::'a,Y) & defined(Y::'a,Z) & identity_map(Y) --> defined(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   613
\  (\\<forall>X. identity_map(domain(X))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   614
\  (\\<forall>X. identity_map(codomain(X))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   615
\  (\\<forall>X. defined(X::'a,domain(X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   616
\  (\\<forall>X. defined(codomain(X),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   617
\  (\\<forall>X. product(X::'a,domain(X),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   618
\  (\\<forall>X. product(codomain(X),X,X)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   619
\  (\\<forall>X Y. defined(X::'a,Y) & identity_map(X) --> product(X::'a,Y,Y)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   620
\  (\\<forall>Y X. defined(X::'a,Y) & identity_map(Y) --> product(X::'a,Y,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   621
\  (\\<forall>X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   622
\  (\\<forall>X Y Z W. equal(X::'a,Y) & product(X::'a,Z,W) --> product(Y::'a,Z,W)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   623
\  (\\<forall>X Z Y W. equal(X::'a,Y) & product(Z::'a,X,W) --> product(Z::'a,Y,W)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   624
\  (\\<forall>X Z W Y. equal(X::'a,Y) & product(Z::'a,W,X) --> product(Z::'a,W,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   625
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(domain(X),domain(Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   626
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(codomain(X),codomain(Y))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   627
\  (\\<forall>X Y. equal(X::'a,Y) & identity_map(X) --> identity_map(Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   628
\  (\\<forall>X Y Z. equal(X::'a,Y) & defined(X::'a,Z) --> defined(Y::'a,Z)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   629
\  (\\<forall>X Z Y. equal(X::'a,Y) & defined(Z::'a,X) --> defined(Z::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   630
\  (\\<forall>X Z Y. equal(X::'a,Y) --> equal(compos(Z::'a,X),compos(Z::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   631
\  (\\<forall>X Y Z. equal(X::'a,Y) --> equal(compos(X::'a,Z),compos(Y::'a,Z))) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   632
\  (defined(a::'a,d)) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   633
\  (identity_map(d)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   634
\  (~equal(domain(a),d)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   635
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   636
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   637
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   638
(*1728 inferences so far.  Searching to depth 10.  5.8 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   639
val CAT007_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   640
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   641
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   642
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   643
\  (\\<forall>X Y. defined(X::'a,Y) --> product(X::'a,Y,compos(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   644
\  (\\<forall>Z X Y. product(X::'a,Y,Z) --> defined(X::'a,Y)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   645
\  (\\<forall>X Xy Y Z. product(X::'a,Y,Xy) & defined(Xy::'a,Z) --> defined(Y::'a,Z)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   646
\  (\\<forall>Y Xy Z X Yz. product(X::'a,Y,Xy) & product(Y::'a,Z,Yz) & defined(Xy::'a,Z) --> defined(X::'a,Yz)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   647
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   648
\  (\\<forall>Z Yz X Y. product(Y::'a,Z,Yz) & defined(X::'a,Yz) --> defined(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   649
\  (\\<forall>Y X Yz Xy Z. product(Y::'a,Z,Yz) & product(X::'a,Y,Xy) & defined(X::'a,Yz) --> defined(Xy::'a,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   650
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   651
\  (\\<forall>Y X Z. defined(X::'a,Y) & defined(Y::'a,Z) & identity_map(Y) --> defined(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   652
\  (\\<forall>X. identity_map(domain(X))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   653
\  (\\<forall>X. identity_map(codomain(X))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   654
\  (\\<forall>X. defined(X::'a,domain(X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   655
\  (\\<forall>X. defined(codomain(X),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   656
\  (\\<forall>X. product(X::'a,domain(X),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   657
\  (\\<forall>X. product(codomain(X),X,X)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   658
\  (\\<forall>X Y. defined(X::'a,Y) & identity_map(X) --> product(X::'a,Y,Y)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   659
\  (\\<forall>Y X. defined(X::'a,Y) & identity_map(Y) --> product(X::'a,Y,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   660
\  (\\<forall>X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   661
\  (\\<forall>X Y Z W. equal(X::'a,Y) & product(X::'a,Z,W) --> product(Y::'a,Z,W)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   662
\  (\\<forall>X Z Y W. equal(X::'a,Y) & product(Z::'a,X,W) --> product(Z::'a,Y,W)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   663
\  (\\<forall>X Z W Y. equal(X::'a,Y) & product(Z::'a,W,X) --> product(Z::'a,W,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   664
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(domain(X),domain(Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   665
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(codomain(X),codomain(Y))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   666
\  (\\<forall>X Y. equal(X::'a,Y) & identity_map(X) --> identity_map(Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   667
\  (\\<forall>X Y Z. equal(X::'a,Y) & defined(X::'a,Z) --> defined(Y::'a,Z)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   668
\  (\\<forall>X Z Y. equal(X::'a,Y) & defined(Z::'a,X) --> defined(Z::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   669
\  (\\<forall>X Z Y. equal(X::'a,Y) --> equal(compos(Z::'a,X),compos(Z::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   670
\  (\\<forall>X Y Z. equal(X::'a,Y) --> equal(compos(X::'a,Z),compos(Y::'a,Z))) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   671
\  (equal(domain(a),codomain(b))) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   672
\  (~defined(a::'a,b)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   673
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   674
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   675
(*82895 inferences so far.  Searching to depth 13.  355 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   676
val CAT018_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   677
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   678
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   679
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   680
\  (\\<forall>X Y. defined(X::'a,Y) --> product(X::'a,Y,compos(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   681
\  (\\<forall>Z X Y. product(X::'a,Y,Z) --> defined(X::'a,Y)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   682
\  (\\<forall>X Xy Y Z. product(X::'a,Y,Xy) & defined(Xy::'a,Z) --> defined(Y::'a,Z)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   683
\  (\\<forall>Y Xy Z X Yz. product(X::'a,Y,Xy) & product(Y::'a,Z,Yz) & defined(Xy::'a,Z) --> defined(X::'a,Yz)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   684
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   685
\  (\\<forall>Z Yz X Y. product(Y::'a,Z,Yz) & defined(X::'a,Yz) --> defined(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   686
\  (\\<forall>Y X Yz Xy Z. product(Y::'a,Z,Yz) & product(X::'a,Y,Xy) & defined(X::'a,Yz) --> defined(Xy::'a,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   687
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   688
\  (\\<forall>Y X Z. defined(X::'a,Y) & defined(Y::'a,Z) & identity_map(Y) --> defined(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   689
\  (\\<forall>X. identity_map(domain(X))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   690
\  (\\<forall>X. identity_map(codomain(X))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   691
\  (\\<forall>X. defined(X::'a,domain(X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   692
\  (\\<forall>X. defined(codomain(X),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   693
\  (\\<forall>X. product(X::'a,domain(X),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   694
\  (\\<forall>X. product(codomain(X),X,X)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   695
\  (\\<forall>X Y. defined(X::'a,Y) & identity_map(X) --> product(X::'a,Y,Y)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   696
\  (\\<forall>Y X. defined(X::'a,Y) & identity_map(Y) --> product(X::'a,Y,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   697
\  (\\<forall>X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   698
\  (\\<forall>X Y Z W. equal(X::'a,Y) & product(X::'a,Z,W) --> product(Y::'a,Z,W)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   699
\  (\\<forall>X Z Y W. equal(X::'a,Y) & product(Z::'a,X,W) --> product(Z::'a,Y,W)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   700
\  (\\<forall>X Z W Y. equal(X::'a,Y) & product(Z::'a,W,X) --> product(Z::'a,W,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   701
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(domain(X),domain(Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   702
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(codomain(X),codomain(Y))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   703
\  (\\<forall>X Y. equal(X::'a,Y) & identity_map(X) --> identity_map(Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   704
\  (\\<forall>X Y Z. equal(X::'a,Y) & defined(X::'a,Z) --> defined(Y::'a,Z)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   705
\  (\\<forall>X Z Y. equal(X::'a,Y) & defined(Z::'a,X) --> defined(Z::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   706
\  (\\<forall>X Z Y. equal(X::'a,Y) --> equal(compos(Z::'a,X),compos(Z::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   707
\  (\\<forall>X Y Z. equal(X::'a,Y) --> equal(compos(X::'a,Z),compos(Y::'a,Z))) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   708
\  (defined(a::'a,b)) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   709
\  (defined(b::'a,c)) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   710
\  (~defined(a::'a,compos(b::'a,c))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   711
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   712
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   713
(*1118 inferences so far.  Searching to depth 8.  2.3 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   714
val COL001_2 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   715
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   716
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   717
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   718
\  (\\<forall>X Y Z. equal(apply(apply(apply(s::'a,X),Y),Z),apply(apply(X::'a,Z),apply(Y::'a,Z)))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   719
\  (\\<forall>Y X. equal(apply(apply(k::'a,X),Y),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   720
\  (\\<forall>X Y Z. equal(apply(apply(apply(b::'a,X),Y),Z),apply(X::'a,apply(Y::'a,Z)))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   721
\  (\\<forall>X. equal(apply(i::'a,X),X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   722
\  (\\<forall>A B C. equal(A::'a,B) --> equal(apply(A::'a,C),apply(B::'a,C))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   723
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(apply(F'::'a,D),apply(F'::'a,E))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   724
\  (\\<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))))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   725
\  (\\<forall>Y. ~equal(Y::'a,apply(combinator::'a,Y))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   726
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   727
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   728
(*500 inferences so far.  Searching to depth 8.  0.9 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   729
val COL023_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   730
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   731
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   732
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   733
\  (\\<forall>X Y Z. equal(apply(apply(apply(b::'a,X),Y),Z),apply(X::'a,apply(Y::'a,Z)))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   734
\  (\\<forall>X Y Z. equal(apply(apply(apply(n::'a,X),Y),Z),apply(apply(apply(X::'a,Z),Y),Z))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   735
\  (\\<forall>A B C. equal(A::'a,B) --> equal(apply(A::'a,C),apply(B::'a,C))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   736
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(apply(F'::'a,D),apply(F'::'a,E))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   737
\  (\\<forall>Y. ~equal(Y::'a,apply(combinator::'a,Y))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   738
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   739
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   740
(*3018 inferences so far.  Searching to depth 10.  4.3 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   741
val COL032_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   742
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   743
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   744
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   745
\  (\\<forall>X. equal(apply(m::'a,X),apply(X::'a,X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   746
\  (\\<forall>Y X Z. equal(apply(apply(apply(q::'a,X),Y),Z),apply(Y::'a,apply(X::'a,Z)))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   747
\  (\\<forall>A B C. equal(A::'a,B) --> equal(apply(A::'a,C),apply(B::'a,C))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   748
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(apply(F'::'a,D),apply(F'::'a,E))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   749
\  (\\<forall>G H. equal(G::'a,H) --> equal(f(G),f(H))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   750
\  (\\<forall>Y. ~equal(apply(Y::'a,f(Y)),apply(f(Y),apply(Y::'a,f(Y))))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   751
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   752
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   753
(*381878 inferences so far.  Searching to depth 13.  670.4 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   754
val COL052_2 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   755
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   756
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   757
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   758
\  (\\<forall>X Y W. equal(response(compos(X::'a,Y),W),response(X::'a,response(Y::'a,W)))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   759
\  (\\<forall>X Y. agreeable(X) --> equal(response(X::'a,common_bird(Y)),response(Y::'a,common_bird(Y)))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   760
\  (\\<forall>Z X. equal(response(X::'a,Z),response(compatible(X),Z)) --> agreeable(X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   761
\  (\\<forall>A B. equal(A::'a,B) --> equal(common_bird(A),common_bird(B))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   762
\  (\\<forall>C D. equal(C::'a,D) --> equal(compatible(C),compatible(D))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   763
\  (\\<forall>Q R. equal(Q::'a,R) & agreeable(Q) --> agreeable(R)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   764
\  (\\<forall>A B C. equal(A::'a,B) --> equal(compos(A::'a,C),compos(B::'a,C))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   765
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(compos(F'::'a,D),compos(F'::'a,E))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   766
\  (\\<forall>G H I'. equal(G::'a,H) --> equal(response(G::'a,I'),response(H::'a,I'))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   767
\  (\\<forall>J L K'. equal(J::'a,K') --> equal(response(L::'a,J),response(L::'a,K'))) &    \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   768
\  (agreeable(c)) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   769
\  (~agreeable(a)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   770
\  (equal(c::'a,compos(a::'a,b))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   771
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   772
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   773
(*13201 inferences so far.  Searching to depth 11.  31.9 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   774
val COL075_2 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   775
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   776
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   777
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   778
\  (\\<forall>Y X. equal(apply(apply(k::'a,X),Y),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   779
\  (\\<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)))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   780
\  (\\<forall>D E F'. equal(D::'a,E) --> equal(apply(D::'a,F'),apply(E::'a,F'))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   781
\  (\\<forall>G I' H. equal(G::'a,H) --> equal(apply(I'::'a,G),apply(I'::'a,H))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   782
\  (\\<forall>A B. equal(A::'a,B) --> equal(b(A),b(B))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   783
\  (\\<forall>C D. equal(C::'a,D) --> equal(c(C),c(D))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   784
\  (\\<forall>Y. ~equal(apply(apply(Y::'a,b(Y)),c(Y)),apply(b(Y),b(Y)))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   785
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   786
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   787
(*33 inferences so far.  Searching to depth 7.  0.1 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   788
val COM001_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   789
 ("(\\<forall>Goal_state Start_state. follows(Goal_state::'a,Start_state) --> succeeds(Goal_state::'a,Start_state)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   790
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   791
\  (\\<forall>Start_state Label Goal_state. has(Start_state::'a,goto(Label)) & labels(Label::'a,Goal_state) --> succeeds(Goal_state::'a,Start_state)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   792
\  (\\<forall>Start_state Condition Goal_state. has(Start_state::'a,ifthen(Condition::'a,Goal_state)) --> succeeds(Goal_state::'a,Start_state)) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   793
\  (labels(loop::'a,p3)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   794
\  (has(p3::'a,ifthen(equal(register_j::'a,n),p4))) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   795
\  (has(p4::'a,goto(out))) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   796
\  (follows(p5::'a,p4)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   797
\  (follows(p8::'a,p3)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   798
\  (has(p8::'a,goto(loop))) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   799
\  (~succeeds(p3::'a,p3)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   800
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   801
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   802
(*533 inferences so far.  Searching to depth 13.  0.3 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   803
val COM002_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   804
 ("(\\<forall>Goal_state Start_state. follows(Goal_state::'a,Start_state) --> succeeds(Goal_state::'a,Start_state)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   805
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   806
\  (\\<forall>Start_state Label Goal_state. has(Start_state::'a,goto(Label)) & labels(Label::'a,Goal_state) --> succeeds(Goal_state::'a,Start_state)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   807
\  (\\<forall>Start_state Condition Goal_state. has(Start_state::'a,ifthen(Condition::'a,Goal_state)) --> succeeds(Goal_state::'a,Start_state)) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   808
\  (has(p1::'a,assign(register_j::'a,num0))) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   809
\  (follows(p2::'a,p1)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   810
\  (has(p2::'a,assign(register_k::'a,num1))) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   811
\  (labels(loop::'a,p3)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   812
\  (follows(p3::'a,p2)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   813
\  (has(p3::'a,ifthen(equal(register_j::'a,n),p4))) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   814
\  (has(p4::'a,goto(out))) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   815
\  (follows(p5::'a,p4)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   816
\  (follows(p6::'a,p3)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   817
\  (has(p6::'a,assign(register_k::'a,times(num2::'a,register_k)))) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   818
\  (follows(p7::'a,p6)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   819
\  (has(p7::'a,assign(register_j::'a,plus(register_j::'a,num1)))) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   820
\  (follows(p8::'a,p7)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   821
\  (has(p8::'a,goto(loop))) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   822
\  (~succeeds(p3::'a,p3)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   823
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   824
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   825
(*4821 inferences so far.  Searching to depth 14.  1.3 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   826
val COM002_2 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   827
 ("(\\<forall>Goal_state Start_state. ~(fails(Goal_state::'a,Start_state) & follows(Goal_state::'a,Start_state))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   828
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   829
\  (\\<forall>Start_state Label Goal_state. ~(fails(Goal_state::'a,Start_state) & has(Start_state::'a,goto(Label)) & labels(Label::'a,Goal_state))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   830
\  (\\<forall>Start_state Condition Goal_state. ~(fails(Goal_state::'a,Start_state) & has(Start_state::'a,ifthen(Condition::'a,Goal_state)))) &     \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   831
\  (has(p1::'a,assign(register_j::'a,num0))) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   832
\  (follows(p2::'a,p1)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   833
\  (has(p2::'a,assign(register_k::'a,num1))) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   834
\  (labels(loop::'a,p3)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   835
\  (follows(p3::'a,p2)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   836
\  (has(p3::'a,ifthen(equal(register_j::'a,n),p4))) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   837
\  (has(p4::'a,goto(out))) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   838
\  (follows(p5::'a,p4)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   839
\  (follows(p6::'a,p3)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   840
\  (has(p6::'a,assign(register_k::'a,times(num2::'a,register_k)))) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   841
\  (follows(p7::'a,p6)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   842
\  (has(p7::'a,assign(register_j::'a,plus(register_j::'a,num1)))) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   843
\  (follows(p8::'a,p7)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   844
\  (has(p8::'a,goto(loop))) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   845
\  (fails(p3::'a,p3)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   846
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   847
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   848
(*98 inferences so far.  Searching to depth 10.  1.1 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   849
val COM003_2 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   850
 ("(\\<forall>X Y Z. program_decides(X) & program(Y) --> decides(X::'a,Y,Z)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   851
\  (\\<forall>X. program_decides(X) | program(f2(X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   852
\  (\\<forall>X. decides(X::'a,f2(X),f1(X)) --> program_decides(X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   853
\  (\\<forall>X. program_program_decides(X) --> program(X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   854
\  (\\<forall>X. program_program_decides(X) --> program_decides(X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   855
\  (\\<forall>X. program(X) & program_decides(X) --> program_program_decides(X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   856
\  (\\<forall>X. algorithm_program_decides(X) --> algorithm(X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   857
\  (\\<forall>X. algorithm_program_decides(X) --> program_decides(X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   858
\  (\\<forall>X. algorithm(X) & program_decides(X) --> algorithm_program_decides(X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   859
\  (\\<forall>Y X. program_halts2(X::'a,Y) --> program(X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   860
\  (\\<forall>X Y. program_halts2(X::'a,Y) --> halts2(X::'a,Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   861
\  (\\<forall>X Y. program(X) & halts2(X::'a,Y) --> program_halts2(X::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   862
\  (\\<forall>W X Y Z. halts3_outputs(X::'a,Y,Z,W) --> halts3(X::'a,Y,Z)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   863
\  (\\<forall>Y Z X W. halts3_outputs(X::'a,Y,Z,W) --> outputs(X::'a,W)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   864
\  (\\<forall>Y Z X W. halts3(X::'a,Y,Z) & outputs(X::'a,W) --> halts3_outputs(X::'a,Y,Z,W)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   865
\  (\\<forall>Y X. program_not_halts2(X::'a,Y) --> program(X)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   866
\  (\\<forall>X Y. ~(program_not_halts2(X::'a,Y) & halts2(X::'a,Y))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   867
\  (\\<forall>X Y. program(X) --> program_not_halts2(X::'a,Y) | halts2(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   868
\  (\\<forall>W X Y. halts2_outputs(X::'a,Y,W) --> halts2(X::'a,Y)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   869
\  (\\<forall>Y X W. halts2_outputs(X::'a,Y,W) --> outputs(X::'a,W)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   870
\  (\\<forall>Y X W. halts2(X::'a,Y) & outputs(X::'a,W) --> halts2_outputs(X::'a,Y,W)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   871
\  (\\<forall>X W Y Z. program_halts2_halts3_outputs(X::'a,Y,Z,W) --> program_halts2(Y::'a,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   872
\  (\\<forall>X Y Z W. program_halts2_halts3_outputs(X::'a,Y,Z,W) --> halts3_outputs(X::'a,Y,Z,W)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   873
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   874
\  (\\<forall>X W Y Z. program_not_halts2_halts3_outputs(X::'a,Y,Z,W) --> program_not_halts2(Y::'a,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   875
\  (\\<forall>X Y Z W. program_not_halts2_halts3_outputs(X::'a,Y,Z,W) --> halts3_outputs(X::'a,Y,Z,W)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   876
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   877
\  (\\<forall>X W Y. program_halts2_halts2_outputs(X::'a,Y,W) --> program_halts2(Y::'a,Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   878
\  (\\<forall>X Y W. program_halts2_halts2_outputs(X::'a,Y,W) --> halts2_outputs(X::'a,Y,W)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   879
\  (\\<forall>X Y W. program_halts2(Y::'a,Y) & halts2_outputs(X::'a,Y,W) --> program_halts2_halts2_outputs(X::'a,Y,W)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   880
\  (\\<forall>X W Y. program_not_halts2_halts2_outputs(X::'a,Y,W) --> program_not_halts2(Y::'a,Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   881
\  (\\<forall>X Y W. program_not_halts2_halts2_outputs(X::'a,Y,W) --> halts2_outputs(X::'a,Y,W)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   882
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   883
\  (\\<forall>X. algorithm_program_decides(X) --> program_program_decides(c1)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   884
\  (\\<forall>W Y Z. program_program_decides(W) --> program_halts2_halts3_outputs(W::'a,Y,Z,good)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   885
\  (\\<forall>W Y Z. program_program_decides(W) --> program_not_halts2_halts3_outputs(W::'a,Y,Z,bad)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   886
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   887
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   888
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   889
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   890
\  (\\<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)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   891
\  (\\<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)) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   892
\  (algorithm_program_decides(c4)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   893
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   894
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   895
(****************SLOW
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   896
2100398 inferences so far.  Searching to depth 12.  No proof after 30 mins.
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   897
val COM004_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   898
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   899
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   900
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   901
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   902
\  (\\<forall>X. contradictory(negate(X),X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   903
\  (\\<forall>X. contradictory(X::'a,negate(X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   904
\  (\\<forall>X. siblings(left_child_of(X),right_child_of(X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   905
\  (\\<forall>D E. equal(D::'a,E) --> equal(left_child_of(D),left_child_of(E))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   906
\  (\\<forall>F' G. equal(F'::'a,G) --> equal(negate(F'),negate(G))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   907
\  (\\<forall>H I' J. equal(H::'a,I') --> equal(or(H::'a,J),or(I'::'a,J))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   908
\  (\\<forall>K' M L. equal(K'::'a,L) --> equal(or(M::'a,K'),or(M::'a,L))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   909
\  (\\<forall>N O_ P. equal(N::'a,O_) --> equal(parent_of(N::'a,P),parent_of(O_::'a,P))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   910
\  (\\<forall>Q S' R. equal(Q::'a,R) --> equal(parent_of(S'::'a,Q),parent_of(S'::'a,R))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   911
\  (\\<forall>T' U. equal(T'::'a,U) --> equal(right_child_of(T'),right_child_of(U))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   912
\  (\\<forall>V W X. equal(V::'a,W) & contradictory(V::'a,X) --> contradictory(W::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   913
\  (\\<forall>Y A1 Z. equal(Y::'a,Z) & contradictory(A1::'a,Y) --> contradictory(A1::'a,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   914
\  (\\<forall>B1 C1 D1. equal(B1::'a,C1) & failure_node(B1::'a,D1) --> failure_node(C1::'a,D1)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   915
\  (\\<forall>E1 G1 F1. equal(E1::'a,F1) & failure_node(G1::'a,E1) --> failure_node(G1::'a,F1)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   916
\  (\\<forall>H1 I1 J1. equal(H1::'a,I1) & siblings(H1::'a,J1) --> siblings(I1::'a,J1)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   917
\  (\\<forall>K1 M1 L1. equal(K1::'a,L1) & siblings(M1::'a,K1) --> siblings(M1::'a,L1)) &   \
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
   918
\  (failure_node(n_left::'a,or(EMPTY::'a,atom))) &      \
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
   919
\  (failure_node(n_right::'a,or(EMPTY::'a,negate(atom)))) &     \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   920
\  (equal(n_left::'a,left_child_of(n))) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   921
\  (equal(n_right::'a,right_child_of(n))) & \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   922
\  (\\<forall>Z. ~failure_node(Z::'a,or(EMPTY::'a,EMPTY))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   923
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   924
****************)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   925
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   926
(*179 inferences so far.  Searching to depth 7.  3.9 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   927
val GEO003_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   928
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   929
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   930
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   931
\  (\\<forall>X Y. between(X::'a,Y,X) --> equal(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   932
\  (\\<forall>V X Y Z. between(X::'a,Y,V) & between(Y::'a,Z,V) --> between(X::'a,Y,Z)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   933
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   934
\  (\\<forall>Y X. equidistant(X::'a,Y,Y,X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   935
\  (\\<forall>Z X Y. equidistant(X::'a,Y,Z,Z) --> equal(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   936
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   937
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   938
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   939
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   940
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   941
\  (\\<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))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   942
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   943
\  (\\<forall>X Y W V. between(X::'a,Y,extension(X::'a,Y,W,V))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   944
\  (\\<forall>X Y W V. equidistant(Y::'a,extension(X::'a,Y,W,V),W,V)) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   945
\  (~between(lower_dimension_point_1::'a,lower_dimension_point_2,lower_dimension_point_3)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   946
\  (~between(lower_dimension_point_2::'a,lower_dimension_point_3,lower_dimension_point_1)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   947
\  (~between(lower_dimension_point_3::'a,lower_dimension_point_1,lower_dimension_point_2)) &        \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   948
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   949
\  (\\<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))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   950
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   951
\  (\\<forall>X Y W Z. equal(X::'a,Y) & between(X::'a,W,Z) --> between(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   952
\  (\\<forall>X W Y Z. equal(X::'a,Y) & between(W::'a,X,Z) --> between(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   953
\  (\\<forall>X W Z Y. equal(X::'a,Y) & between(W::'a,Z,X) --> between(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   954
\  (\\<forall>X Y V W Z. equal(X::'a,Y) & equidistant(X::'a,V,W,Z) --> equidistant(Y::'a,V,W,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   955
\  (\\<forall>X V Y W Z. equal(X::'a,Y) & equidistant(V::'a,X,W,Z) --> equidistant(V::'a,Y,W,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   956
\  (\\<forall>X V W Y Z. equal(X::'a,Y) & equidistant(V::'a,W,X,Z) --> equidistant(V::'a,W,Y,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   957
\  (\\<forall>X V W Z Y. equal(X::'a,Y) & equidistant(V::'a,W,Z,X) --> equidistant(V::'a,W,Z,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   958
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   959
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   960
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   961
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   962
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   963
\  (\\<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'))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   964
\  (\\<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))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   965
\  (\\<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))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   966
\  (\\<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))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   967
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   968
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   969
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   970
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   971
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   972
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   973
\  (\\<forall>X Y V1 V2 V3. equal(X::'a,Y) --> equal(extension(X::'a,V1,V2,V3),extension(Y::'a,V1,V2,V3))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   974
\  (\\<forall>X V1 Y V2 V3. equal(X::'a,Y) --> equal(extension(V1::'a,X,V2,V3),extension(V1::'a,Y,V2,V3))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   975
\  (\\<forall>X V1 V2 Y V3. equal(X::'a,Y) --> equal(extension(V1::'a,V2,X,V3),extension(V1::'a,V2,Y,V3))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   976
\  (\\<forall>X V1 V2 V3 Y. equal(X::'a,Y) --> equal(extension(V1::'a,V2,V3,X),extension(V1::'a,V2,V3,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   977
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   978
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   979
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   980
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   981
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   982
\  (\\<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))) &    \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   983
\  (~between(a::'a,b,b)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
   984
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   985
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   986
(*4272 inferences so far.  Searching to depth 10.  29.4 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
   987
val GEO017_2 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   988
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   989
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   990
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   991
\  (\\<forall>Y X. equidistant(X::'a,Y,Y,X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   992
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   993
\  (\\<forall>Z X Y. equidistant(X::'a,Y,Z,Z) --> equal(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   994
\  (\\<forall>X Y W V. between(X::'a,Y,extension(X::'a,Y,W,V))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   995
\  (\\<forall>X Y W V. equidistant(Y::'a,extension(X::'a,Y,W,V),W,V)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   996
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   997
\  (\\<forall>X Y. between(X::'a,Y,X) --> equal(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   998
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
   999
\  (\\<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)) &     \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1000
\  (~between(lower_dimension_point_1::'a,lower_dimension_point_2,lower_dimension_point_3)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1001
\  (~between(lower_dimension_point_2::'a,lower_dimension_point_3,lower_dimension_point_1)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1002
\  (~between(lower_dimension_point_3::'a,lower_dimension_point_1,lower_dimension_point_2)) &        \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1003
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1004
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1005
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1006
\  (\\<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))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1007
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1008
\  (\\<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))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1009
\  (\\<forall>X Y W Z. equal(X::'a,Y) & between(X::'a,W,Z) --> between(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1010
\  (\\<forall>X W Y Z. equal(X::'a,Y) & between(W::'a,X,Z) --> between(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1011
\  (\\<forall>X W Z Y. equal(X::'a,Y) & between(W::'a,Z,X) --> between(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1012
\  (\\<forall>X Y V W Z. equal(X::'a,Y) & equidistant(X::'a,V,W,Z) --> equidistant(Y::'a,V,W,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1013
\  (\\<forall>X V Y W Z. equal(X::'a,Y) & equidistant(V::'a,X,W,Z) --> equidistant(V::'a,Y,W,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1014
\  (\\<forall>X V W Y Z. equal(X::'a,Y) & equidistant(V::'a,W,X,Z) --> equidistant(V::'a,W,Y,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1015
\  (\\<forall>X V W Z Y. equal(X::'a,Y) & equidistant(V::'a,W,Z,X) --> equidistant(V::'a,W,Z,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1016
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1017
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1018
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1019
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1020
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1021
\  (\\<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'))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1022
\  (\\<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))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1023
\  (\\<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))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1024
\  (\\<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))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1025
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1026
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1027
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1028
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1029
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1030
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1031
\  (\\<forall>X Y V1 V2 V3. equal(X::'a,Y) --> equal(extension(X::'a,V1,V2,V3),extension(Y::'a,V1,V2,V3))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1032
\  (\\<forall>X V1 Y V2 V3. equal(X::'a,Y) --> equal(extension(V1::'a,X,V2,V3),extension(V1::'a,Y,V2,V3))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1033
\  (\\<forall>X V1 V2 Y V3. equal(X::'a,Y) --> equal(extension(V1::'a,V2,X,V3),extension(V1::'a,V2,Y,V3))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1034
\  (\\<forall>X V1 V2 V3 Y. equal(X::'a,Y) --> equal(extension(V1::'a,V2,V3,X),extension(V1::'a,V2,V3,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1035
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1036
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1037
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1038
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1039
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1040
\  (\\<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))) &    \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1041
\  (equidistant(u::'a,v,w,x)) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1042
\  (~equidistant(u::'a,v,x,w)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1043
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1044
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1045
(****************SLOW
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1046
382903 inferences so far.  Searching to depth 9.  No proof after 35 minutes.
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1047
val GEO027_3 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1048
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1049
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1050
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1051
\  (\\<forall>Y X. equidistant(X::'a,Y,Y,X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1052
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1053
\  (\\<forall>Z X Y. equidistant(X::'a,Y,Z,Z) --> equal(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1054
\  (\\<forall>X Y W V. between(X::'a,Y,extension(X::'a,Y,W,V))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1055
\  (\\<forall>X Y W V. equidistant(Y::'a,extension(X::'a,Y,W,V),W,V)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1056
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1057
\  (\\<forall>X Y. between(X::'a,Y,X) --> equal(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1058
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1059
\  (\\<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)) &     \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1060
\  (~between(lower_dimension_point_1::'a,lower_dimension_point_2,lower_dimension_point_3)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1061
\  (~between(lower_dimension_point_2::'a,lower_dimension_point_3,lower_dimension_point_1)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1062
\  (~between(lower_dimension_point_3::'a,lower_dimension_point_1,lower_dimension_point_2)) &        \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1063
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1064
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1065
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1066
\  (\\<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))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1067
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1068
\  (\\<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))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1069
\  (\\<forall>X Y W Z. equal(X::'a,Y) & between(X::'a,W,Z) --> between(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1070
\  (\\<forall>X W Y Z. equal(X::'a,Y) & between(W::'a,X,Z) --> between(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1071
\  (\\<forall>X W Z Y. equal(X::'a,Y) & between(W::'a,Z,X) --> between(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1072
\  (\\<forall>X Y V W Z. equal(X::'a,Y) & equidistant(X::'a,V,W,Z) --> equidistant(Y::'a,V,W,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1073
\  (\\<forall>X V Y W Z. equal(X::'a,Y) & equidistant(V::'a,X,W,Z) --> equidistant(V::'a,Y,W,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1074
\  (\\<forall>X V W Y Z. equal(X::'a,Y) & equidistant(V::'a,W,X,Z) --> equidistant(V::'a,W,Y,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1075
\  (\\<forall>X V W Z Y. equal(X::'a,Y) & equidistant(V::'a,W,Z,X) --> equidistant(V::'a,W,Z,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1076
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1077
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1078
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1079
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1080
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1081
\  (\\<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'))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1082
\  (\\<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))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1083
\  (\\<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))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1084
\  (\\<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))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1085
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1086
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1087
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1088
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1089
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1090
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1091
\  (\\<forall>X Y V1 V2 V3. equal(X::'a,Y) --> equal(extension(X::'a,V1,V2,V3),extension(Y::'a,V1,V2,V3))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1092
\  (\\<forall>X V1 Y V2 V3. equal(X::'a,Y) --> equal(extension(V1::'a,X,V2,V3),extension(V1::'a,Y,V2,V3))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1093
\  (\\<forall>X V1 V2 Y V3. equal(X::'a,Y) --> equal(extension(V1::'a,V2,X,V3),extension(V1::'a,V2,Y,V3))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1094
\  (\\<forall>X V1 V2 V3 Y. equal(X::'a,Y) --> equal(extension(V1::'a,V2,V3,X),extension(V1::'a,V2,V3,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1095
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1096
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1097
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1098
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1099
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1100
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1101
\  (\\<forall>U V. equal(reflection(U::'a,V),extension(U::'a,V,U,V))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1102
\  (\\<forall>X Y Z. equal(X::'a,Y) --> equal(reflection(X::'a,Z),reflection(Y::'a,Z))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1103
\  (\\<forall>A1 C1 B1. equal(A1::'a,B1) --> equal(reflection(C1::'a,A1),reflection(C1::'a,B1))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1104
\  (\\<forall>U V. equidistant(U::'a,V,U,V)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1105
\  (\\<forall>W X U V. equidistant(U::'a,V,W,X) --> equidistant(W::'a,X,U,V)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1106
\  (\\<forall>V U W X. equidistant(U::'a,V,W,X) --> equidistant(V::'a,U,W,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1107
\  (\\<forall>U V X W. equidistant(U::'a,V,W,X) --> equidistant(U::'a,V,X,W)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1108
\  (\\<forall>V U X W. equidistant(U::'a,V,W,X) --> equidistant(V::'a,U,X,W)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1109
\  (\\<forall>W X V U. equidistant(U::'a,V,W,X) --> equidistant(W::'a,X,V,U)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1110
\  (\\<forall>X W U V. equidistant(U::'a,V,W,X) --> equidistant(X::'a,W,U,V)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1111
\  (\\<forall>X W V U. equidistant(U::'a,V,W,X) --> equidistant(X::'a,W,V,U)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1112
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1113
\  (\\<forall>U V W. equal(V::'a,extension(U::'a,V,W,W))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1114
\  (\\<forall>W X U V Y. equal(Y::'a,extension(U::'a,V,W,X)) --> between(U::'a,V,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1115
\  (\\<forall>U V. between(U::'a,V,reflection(U::'a,V))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1116
\  (\\<forall>U V. equidistant(V::'a,reflection(U::'a,V),U,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1117
\  (\\<forall>U V. equal(U::'a,V) --> equal(V::'a,reflection(U::'a,V))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1118
\  (\\<forall>U. equal(U::'a,reflection(U::'a,U))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1119
\  (\\<forall>U V. equal(V::'a,reflection(U::'a,V)) --> equal(U::'a,V)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1120
\  (\\<forall>U V. equidistant(U::'a,U,V,V)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1121
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1122
\  (\\<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)) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1123
\  (between(u::'a,v,w)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1124
\  (~equal(u::'a,v)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1125
\  (~equal(w::'a,extension(u::'a,v,v,w))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1126
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1127
****************)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1128
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1129
(*313884 inferences so far.  Searching to depth 10.  887 secs: 15 mins.*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1130
val GEO058_2 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1131
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1132
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1133
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1134
\  (\\<forall>Y X. equidistant(X::'a,Y,Y,X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1135
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1136
\  (\\<forall>Z X Y. equidistant(X::'a,Y,Z,Z) --> equal(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1137
\  (\\<forall>X Y W V. between(X::'a,Y,extension(X::'a,Y,W,V))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1138
\  (\\<forall>X Y W V. equidistant(Y::'a,extension(X::'a,Y,W,V),W,V)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1139
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1140
\  (\\<forall>X Y. between(X::'a,Y,X) --> equal(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1141
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1142
\  (\\<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)) &     \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1143
\  (~between(lower_dimension_point_1::'a,lower_dimension_point_2,lower_dimension_point_3)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1144
\  (~between(lower_dimension_point_2::'a,lower_dimension_point_3,lower_dimension_point_1)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1145
\  (~between(lower_dimension_point_3::'a,lower_dimension_point_1,lower_dimension_point_2)) &        \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1146
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1147
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1148
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1149
\  (\\<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))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1150
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1151
\  (\\<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))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1152
\  (\\<forall>X Y W Z. equal(X::'a,Y) & between(X::'a,W,Z) --> between(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1153
\  (\\<forall>X W Y Z. equal(X::'a,Y) & between(W::'a,X,Z) --> between(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1154
\  (\\<forall>X W Z Y. equal(X::'a,Y) & between(W::'a,Z,X) --> between(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1155
\  (\\<forall>X Y V W Z. equal(X::'a,Y) & equidistant(X::'a,V,W,Z) --> equidistant(Y::'a,V,W,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1156
\  (\\<forall>X V Y W Z. equal(X::'a,Y) & equidistant(V::'a,X,W,Z) --> equidistant(V::'a,Y,W,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1157
\  (\\<forall>X V W Y Z. equal(X::'a,Y) & equidistant(V::'a,W,X,Z) --> equidistant(V::'a,W,Y,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1158
\  (\\<forall>X V W Z Y. equal(X::'a,Y) & equidistant(V::'a,W,Z,X) --> equidistant(V::'a,W,Z,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1159
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1160
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1161
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1162
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1163
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1164
\  (\\<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'))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1165
\  (\\<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))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1166
\  (\\<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))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1167
\  (\\<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))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1168
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1169
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1170
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1171
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1172
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1173
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1174
\  (\\<forall>X Y V1 V2 V3. equal(X::'a,Y) --> equal(extension(X::'a,V1,V2,V3),extension(Y::'a,V1,V2,V3))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1175
\  (\\<forall>X V1 Y V2 V3. equal(X::'a,Y) --> equal(extension(V1::'a,X,V2,V3),extension(V1::'a,Y,V2,V3))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1176
\  (\\<forall>X V1 V2 Y V3. equal(X::'a,Y) --> equal(extension(V1::'a,V2,X,V3),extension(V1::'a,V2,Y,V3))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1177
\  (\\<forall>X V1 V2 V3 Y. equal(X::'a,Y) --> equal(extension(V1::'a,V2,V3,X),extension(V1::'a,V2,V3,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1178
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1179
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1180
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1181
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1182
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1183
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1184
\  (\\<forall>U V. equal(reflection(U::'a,V),extension(U::'a,V,U,V))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1185
\  (\\<forall>X Y Z. equal(X::'a,Y) --> equal(reflection(X::'a,Z),reflection(Y::'a,Z))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1186
\  (\\<forall>A1 C1 B1. equal(A1::'a,B1) --> equal(reflection(C1::'a,A1),reflection(C1::'a,B1))) &  \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1187
\  (equal(v::'a,reflection(u::'a,v))) & \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1188
\  (~equal(u::'a,v)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1189
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1190
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1191
(*0 inferences so far.  Searching to depth 0.  0.2 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1192
val GEO079_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1193
 ("(\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1194
\  (\\<forall>U V W X Y Z. CONGRUENT(U::'a,V,W,X,Y,Z) --> eq(U::'a,V,W,X,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1195
\  (\\<forall>V W U X. trapezoid(U::'a,V,W,X) --> parallel(V::'a,W,U,X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1196
\  (\\<forall>U V X Y. parallel(U::'a,V,X,Y) --> eq(X::'a,V,U,V,X,Y)) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1197
\  (trapezoid(a::'a,b,c,d)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1198
\  (~eq(a::'a,c,b,c,a,d)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1199
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1200
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1201
(****************SLOW
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1202
2032008 inferences so far.  Searching to depth 16.  No proof after 30 minutes.
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1203
val GRP001_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1204
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1205
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1206
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1207
\  (\\<forall>X. product(identity::'a,X,X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1208
\  (\\<forall>X. product(X::'a,identity,X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1209
\  (\\<forall>X. product(INVERSE(X),X,identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1210
\  (\\<forall>X. product(X::'a,INVERSE(X),identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1211
\  (\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1212
\  (\\<forall>X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1213
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1214
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1215
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(INVERSE(X),INVERSE(Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1216
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1217
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1218
\  (\\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1219
\  (\\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1220
\  (\\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1221
\  (\\<forall>X. product(X::'a,X,identity)) &       \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1222
\  (product(a::'a,b,c)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1223
\  (~product(b::'a,a,c)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1224
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1225
****************)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1226
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1227
(*2386 inferences so far.  Searching to depth 11.  8.7 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1228
val GRP008_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1229
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1230
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1231
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1232
\  (\\<forall>X. product(identity::'a,X,X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1233
\  (\\<forall>X. product(X::'a,identity,X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1234
\  (\\<forall>X. product(INVERSE(X),X,identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1235
\  (\\<forall>X. product(X::'a,INVERSE(X),identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1236
\  (\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1237
\  (\\<forall>X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1238
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1239
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1240
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(INVERSE(X),INVERSE(Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1241
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1242
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1243
\  (\\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1244
\  (\\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1245
\  (\\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1246
\  (\\<forall>A B. equal(A::'a,B) --> equal(h(A),h(B))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1247
\  (\\<forall>C D. equal(C::'a,D) --> equal(j(C),j(D))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1248
\  (\\<forall>A B. equal(A::'a,B) & q(A) --> q(B)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1249
\  (\\<forall>B A C. q(A) & product(A::'a,B,C) --> product(B::'a,A,C)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1250
\  (\\<forall>A. product(j(A),A,h(A)) | product(A::'a,j(A),h(A)) | q(A)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1251
\  (\\<forall>A. product(j(A),A,h(A)) & product(A::'a,j(A),h(A)) --> q(A)) &        \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1252
\  (~q(identity)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1253
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1254
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1255
(*8625 inferences so far.  Searching to depth 11.  20 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1256
val GRP013_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1257
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1258
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1259
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1260
\  (\\<forall>X. product(identity::'a,X,X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1261
\  (\\<forall>X. product(X::'a,identity,X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1262
\  (\\<forall>X. product(INVERSE(X),X,identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1263
\  (\\<forall>X. product(X::'a,INVERSE(X),identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1264
\  (\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1265
\  (\\<forall>X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1266
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1267
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1268
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(INVERSE(X),INVERSE(Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1269
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1270
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1271
\  (\\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1272
\  (\\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1273
\  (\\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1274
\  (\\<forall>A. product(A::'a,A,identity)) &       \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1275
\  (product(a::'a,b,c)) &   \
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  1276
\  (product(INVERSE(a),INVERSE(b),d)) & \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1277
\  (\\<forall>A C B. product(INVERSE(A),INVERSE(B),C) --> product(A::'a,C,B)) &     \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1278
\  (~product(c::'a,d,identity)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1279
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1280
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1281
(*2448 inferences so far.  Searching to depth 10.  7.2 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1282
val GRP037_3 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1283
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1284
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1285
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1286
\  (\\<forall>X. product(identity::'a,X,X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1287
\  (\\<forall>X. product(X::'a,identity,X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1288
\  (\\<forall>X. product(INVERSE(X),X,identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1289
\  (\\<forall>X. product(X::'a,INVERSE(X),identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1290
\  (\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1291
\  (\\<forall>X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1292
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1293
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1294
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(INVERSE(X),INVERSE(Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1295
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1296
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1297
\  (\\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1298
\  (\\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1299
\  (\\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1300
\  (\\<forall>A B C. subgroup_member(A) & subgroup_member(B) & product(A::'a,INVERSE(B),C) --> subgroup_member(C)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1301
\  (\\<forall>A B. equal(A::'a,B) & subgroup_member(A) --> subgroup_member(B)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1302
\  (\\<forall>A. subgroup_member(A) --> product(another_identity::'a,A,A)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1303
\  (\\<forall>A. subgroup_member(A) --> product(A::'a,another_identity,A)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1304
\  (\\<forall>A. subgroup_member(A) --> product(A::'a,another_inverse(A),another_identity)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1305
\  (\\<forall>A. subgroup_member(A) --> product(another_inverse(A),A,another_identity)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1306
\  (\\<forall>A. subgroup_member(A) --> subgroup_member(another_inverse(A))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1307
\  (\\<forall>A B. equal(A::'a,B) --> equal(another_inverse(A),another_inverse(B))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1308
\  (\\<forall>A C D B. product(A::'a,B,C) & product(A::'a,D,C) --> equal(D::'a,B)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1309
\  (\\<forall>B C D A. product(A::'a,B,C) & product(D::'a,B,C) --> equal(D::'a,A)) &        \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1310
\  (subgroup_member(a)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1311
\  (subgroup_member(another_identity)) &        \
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  1312
\  (~equal(INVERSE(a),another_inverse(a))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1313
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1314
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1315
(*163 inferences so far.  Searching to depth 11.  0.3 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1316
val GRP031_2 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1317
 ("(\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1318
\  (\\<forall>X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1319
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1320
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1321
\  (\\<forall>A. product(A::'a,INVERSE(A),identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1322
\  (\\<forall>A. product(A::'a,identity,A)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1323
\  (\\<forall>A. ~product(A::'a,a,identity)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1324
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1325
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1326
(*47 inferences so far.  Searching to depth 11.   0.2 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1327
val GRP034_4 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1328
 ("(\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1329
\  (\\<forall>X. product(identity::'a,X,X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1330
\  (\\<forall>X. product(X::'a,identity,X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1331
\  (\\<forall>X. product(X::'a,INVERSE(X),identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1332
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1333
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1334
\  (\\<forall>B A C. subgroup_member(A) & subgroup_member(B) & product(B::'a,INVERSE(A),C) --> subgroup_member(C)) &        \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1335
\  (subgroup_member(a)) &       \
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  1336
\  (~subgroup_member(INVERSE(a))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1337
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1338
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1339
(*3287 inferences so far.  Searching to depth 14.  3.5 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1340
val GRP047_2 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1341
 ("(\\<forall>X. product(identity::'a,X,X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1342
\  (\\<forall>X. product(INVERSE(X),X,identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1343
\  (\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1344
\  (\\<forall>X Y Z W. product(X::'a,Y,Z) & product(X::'a,Y,W) --> equal(Z::'a,W)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1345
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1346
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1347
\  (\\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) &        \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1348
\  (equal(a::'a,b)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1349
\  (~equal(multiply(c::'a,a),multiply(c::'a,b))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1350
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1351
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1352
(*25559 inferences so far.  Searching to depth 19.  16.9 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1353
val GRP130_1_002 = prove_hard
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1354
 ("(group_element(e_1)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1355
\  (group_element(e_2)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1356
\  (~equal(e_1::'a,e_2)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1357
\  (~equal(e_2::'a,e_1)) &  \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1358
\  (\\<forall>X Y. group_element(X) & group_element(Y) --> product(X::'a,Y,e_1) | product(X::'a,Y,e_2)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1359
\  (\\<forall>X Y W Z. product(X::'a,Y,W) & product(X::'a,Y,Z) --> equal(W::'a,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1360
\  (\\<forall>X Y W Z. product(X::'a,W,Y) & product(X::'a,Z,Y) --> equal(W::'a,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1361
\  (\\<forall>Y X W Z. product(W::'a,Y,X) & product(Z::'a,Y,X) --> equal(W::'a,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1362
\  (\\<forall>Z1 Z2 Y X. product(X::'a,Y,Z1) & product(X::'a,Z1,Z2) --> product(Z2::'a,Y,X)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1363
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1364
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1365
(*3468 inferences so far.  Searching to depth 10.  9.1 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1366
val GRP156_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1367
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1368
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1369
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1370
\  (\\<forall>X. equal(multiply(identity::'a,X),X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1371
\  (\\<forall>X. equal(multiply(INVERSE(X),X),identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1372
\  (\\<forall>X Y Z. equal(multiply(multiply(X::'a,Y),Z),multiply(X::'a,multiply(Y::'a,Z)))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1373
\  (\\<forall>A B. equal(A::'a,B) --> equal(INVERSE(A),INVERSE(B))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1374
\  (\\<forall>C D E. equal(C::'a,D) --> equal(multiply(C::'a,E),multiply(D::'a,E))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1375
\  (\\<forall>F' H G. equal(F'::'a,G) --> equal(multiply(H::'a,F'),multiply(H::'a,G))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1376
\  (\\<forall>Y X. equal(greatest_lower_bound(X::'a,Y),greatest_lower_bound(Y::'a,X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1377
\  (\\<forall>Y X. equal(least_upper_bound(X::'a,Y),least_upper_bound(Y::'a,X))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1378
\  (\\<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))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1379
\  (\\<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))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1380
\  (\\<forall>X. equal(least_upper_bound(X::'a,X),X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1381
\  (\\<forall>X. equal(greatest_lower_bound(X::'a,X),X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1382
\  (\\<forall>Y X. equal(least_upper_bound(X::'a,greatest_lower_bound(X::'a,Y)),X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1383
\  (\\<forall>Y X. equal(greatest_lower_bound(X::'a,least_upper_bound(X::'a,Y)),X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1384
\  (\\<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)))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1385
\  (\\<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)))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1386
\  (\\<forall>Y Z X. equal(multiply(least_upper_bound(Y::'a,Z),X),least_upper_bound(multiply(Y::'a,X),multiply(Z::'a,X)))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1387
\  (\\<forall>Y Z X. equal(multiply(greatest_lower_bound(Y::'a,Z),X),greatest_lower_bound(multiply(Y::'a,X),multiply(Z::'a,X)))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1388
\  (\\<forall>A B C. equal(A::'a,B) --> equal(greatest_lower_bound(A::'a,C),greatest_lower_bound(B::'a,C))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1389
\  (\\<forall>A C B. equal(A::'a,B) --> equal(greatest_lower_bound(C::'a,A),greatest_lower_bound(C::'a,B))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1390
\  (\\<forall>A B C. equal(A::'a,B) --> equal(least_upper_bound(A::'a,C),least_upper_bound(B::'a,C))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1391
\  (\\<forall>A C B. equal(A::'a,B) --> equal(least_upper_bound(C::'a,A),least_upper_bound(C::'a,B))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1392
\  (\\<forall>A B C. equal(A::'a,B) --> equal(multiply(A::'a,C),multiply(B::'a,C))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1393
\  (\\<forall>A C B. equal(A::'a,B) --> equal(multiply(C::'a,A),multiply(C::'a,B))) &       \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1394
\  (equal(least_upper_bound(a::'a,b),b)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1395
\  (~equal(greatest_lower_bound(multiply(a::'a,c),multiply(b::'a,c)),multiply(a::'a,c))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1396
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1397
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1398
(*4394 inferences so far.  Searching to depth 10.  8.2 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1399
val GRP168_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1400
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1401
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1402
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1403
\  (\\<forall>X. equal(multiply(identity::'a,X),X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1404
\  (\\<forall>X. equal(multiply(INVERSE(X),X),identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1405
\  (\\<forall>X Y Z. equal(multiply(multiply(X::'a,Y),Z),multiply(X::'a,multiply(Y::'a,Z)))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1406
\  (\\<forall>A B. equal(A::'a,B) --> equal(INVERSE(A),INVERSE(B))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1407
\  (\\<forall>C D E. equal(C::'a,D) --> equal(multiply(C::'a,E),multiply(D::'a,E))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1408
\  (\\<forall>F' H G. equal(F'::'a,G) --> equal(multiply(H::'a,F'),multiply(H::'a,G))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1409
\  (\\<forall>Y X. equal(greatest_lower_bound(X::'a,Y),greatest_lower_bound(Y::'a,X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1410
\  (\\<forall>Y X. equal(least_upper_bound(X::'a,Y),least_upper_bound(Y::'a,X))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1411
\  (\\<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))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1412
\  (\\<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))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1413
\  (\\<forall>X. equal(least_upper_bound(X::'a,X),X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1414
\  (\\<forall>X. equal(greatest_lower_bound(X::'a,X),X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1415
\  (\\<forall>Y X. equal(least_upper_bound(X::'a,greatest_lower_bound(X::'a,Y)),X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1416
\  (\\<forall>Y X. equal(greatest_lower_bound(X::'a,least_upper_bound(X::'a,Y)),X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1417
\  (\\<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)))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1418
\  (\\<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)))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1419
\  (\\<forall>Y Z X. equal(multiply(least_upper_bound(Y::'a,Z),X),least_upper_bound(multiply(Y::'a,X),multiply(Z::'a,X)))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1420
\  (\\<forall>Y Z X. equal(multiply(greatest_lower_bound(Y::'a,Z),X),greatest_lower_bound(multiply(Y::'a,X),multiply(Z::'a,X)))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1421
\  (\\<forall>A B C. equal(A::'a,B) --> equal(greatest_lower_bound(A::'a,C),greatest_lower_bound(B::'a,C))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1422
\  (\\<forall>A C B. equal(A::'a,B) --> equal(greatest_lower_bound(C::'a,A),greatest_lower_bound(C::'a,B))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1423
\  (\\<forall>A B C. equal(A::'a,B) --> equal(least_upper_bound(A::'a,C),least_upper_bound(B::'a,C))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1424
\  (\\<forall>A C B. equal(A::'a,B) --> equal(least_upper_bound(C::'a,A),least_upper_bound(C::'a,B))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1425
\  (\\<forall>A B C. equal(A::'a,B) --> equal(multiply(A::'a,C),multiply(B::'a,C))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1426
\  (\\<forall>A C B. equal(A::'a,B) --> equal(multiply(C::'a,A),multiply(C::'a,B))) &       \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1427
\  (equal(least_upper_bound(a::'a,b),b)) &  \
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  1428
\  (~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",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1429
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1430
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1431
(*250258 inferences so far.  Searching to depth 16.  406.2 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1432
val HEN003_3 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1433
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1434
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1435
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1436
\  (\\<forall>X Y. less_equal(X::'a,Y) --> equal(divide(X::'a,Y),zero)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1437
\  (\\<forall>X Y. equal(divide(X::'a,Y),zero) --> less_equal(X::'a,Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1438
\  (\\<forall>Y X. less_equal(divide(X::'a,Y),X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1439
\  (\\<forall>X Y Z. less_equal(divide(divide(X::'a,Z),divide(Y::'a,Z)),divide(divide(X::'a,Y),Z))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1440
\  (\\<forall>X. less_equal(zero::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1441
\  (\\<forall>X Y. less_equal(X::'a,Y) & less_equal(Y::'a,X) --> equal(X::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1442
\  (\\<forall>X. less_equal(X::'a,identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1443
\  (\\<forall>A B C. equal(A::'a,B) --> equal(divide(A::'a,C),divide(B::'a,C))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1444
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(divide(F'::'a,D),divide(F'::'a,E))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1445
\  (\\<forall>G H I'. equal(G::'a,H) & less_equal(G::'a,I') --> less_equal(H::'a,I')) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1446
\  (\\<forall>J L K'. equal(J::'a,K') & less_equal(L::'a,J) --> less_equal(L::'a,K')) &     \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1447
\  (~equal(divide(a::'a,a),zero)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1448
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1449
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1450
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1451
(*202177 inferences so far.  Searching to depth 14.  451 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1452
val HEN007_2 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1453
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1454
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1455
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1456
\  (\\<forall>X Y. less_equal(X::'a,Y) --> quotient(X::'a,Y,zero)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1457
\  (\\<forall>X Y. quotient(X::'a,Y,zero) --> less_equal(X::'a,Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1458
\  (\\<forall>Y Z X. quotient(X::'a,Y,Z) --> less_equal(Z::'a,X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1459
\  (\\<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) --> less_equal(V4::'a,V5)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1460
\  (\\<forall>X. less_equal(zero::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1461
\  (\\<forall>X Y. less_equal(X::'a,Y) & less_equal(Y::'a,X) --> equal(X::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1462
\  (\\<forall>X. less_equal(X::'a,identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1463
\  (\\<forall>X Y. quotient(X::'a,Y,divide(X::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1464
\  (\\<forall>X Y Z W. quotient(X::'a,Y,Z) & quotient(X::'a,Y,W) --> equal(Z::'a,W)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1465
\  (\\<forall>X Y W Z. equal(X::'a,Y) & quotient(X::'a,W,Z) --> quotient(Y::'a,W,Z)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1466
\  (\\<forall>X W Y Z. equal(X::'a,Y) & quotient(W::'a,X,Z) --> quotient(W::'a,Y,Z)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1467
\  (\\<forall>X W Z Y. equal(X::'a,Y) & quotient(W::'a,Z,X) --> quotient(W::'a,Z,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1468
\  (\\<forall>X Z Y. equal(X::'a,Y) & less_equal(Z::'a,X) --> less_equal(Z::'a,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1469
\  (\\<forall>X Y Z. equal(X::'a,Y) & less_equal(X::'a,Z) --> less_equal(Y::'a,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1470
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(divide(X::'a,W),divide(Y::'a,W))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1471
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(divide(W::'a,X),divide(W::'a,Y))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1472
\  (\\<forall>X. quotient(X::'a,identity,zero)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1473
\  (\\<forall>X. quotient(zero::'a,X,zero)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1474
\  (\\<forall>X. quotient(X::'a,X,zero)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1475
\  (\\<forall>X. quotient(X::'a,zero,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1476
\  (\\<forall>Y X Z. less_equal(X::'a,Y) & less_equal(Y::'a,Z) --> less_equal(X::'a,Z)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1477
\  (\\<forall>W1 X Z W2 Y. quotient(X::'a,Y,W1) & less_equal(W1::'a,Z) & quotient(X::'a,Z,W2) --> less_equal(W2::'a,Y)) &       \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1478
\  (less_equal(x::'a,y)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1479
\  (quotient(z::'a,y,zQy)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1480
\  (quotient(z::'a,x,zQx)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1481
\  (~less_equal(zQy::'a,zQx)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1482
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1483
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1484
(*60026 inferences so far.  Searching to depth 12.  42.2 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1485
val HEN008_4 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1486
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1487
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1488
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1489
\  (\\<forall>X Y. less_equal(X::'a,Y) --> equal(divide(X::'a,Y),zero)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1490
\  (\\<forall>X Y. equal(divide(X::'a,Y),zero) --> less_equal(X::'a,Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1491
\  (\\<forall>Y X. less_equal(divide(X::'a,Y),X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1492
\  (\\<forall>X Y Z. less_equal(divide(divide(X::'a,Z),divide(Y::'a,Z)),divide(divide(X::'a,Y),Z))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1493
\  (\\<forall>X. less_equal(zero::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1494
\  (\\<forall>X Y. less_equal(X::'a,Y) & less_equal(Y::'a,X) --> equal(X::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1495
\  (\\<forall>X. less_equal(X::'a,identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1496
\  (\\<forall>A B C. equal(A::'a,B) --> equal(divide(A::'a,C),divide(B::'a,C))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1497
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(divide(F'::'a,D),divide(F'::'a,E))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1498
\  (\\<forall>G H I'. equal(G::'a,H) & less_equal(G::'a,I') --> less_equal(H::'a,I')) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1499
\  (\\<forall>J L K'. equal(J::'a,K') & less_equal(L::'a,J) --> less_equal(L::'a,K')) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1500
\  (\\<forall>X. equal(divide(X::'a,identity),zero)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1501
\  (\\<forall>X. equal(divide(zero::'a,X),zero)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1502
\  (\\<forall>X. equal(divide(X::'a,X),zero)) &     \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1503
\  (equal(divide(a::'a,zero),a)) &  \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1504
\  (\\<forall>Y X Z. less_equal(X::'a,Y) & less_equal(Y::'a,Z) --> less_equal(X::'a,Z)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1505
\  (\\<forall>X Z Y. less_equal(divide(X::'a,Y),Z) --> less_equal(divide(X::'a,Z),Y)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1506
\  (\\<forall>Y Z X. less_equal(X::'a,Y) --> less_equal(divide(Z::'a,Y),divide(Z::'a,X))) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1507
\  (less_equal(a::'a,b)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1508
\  (~less_equal(divide(a::'a,c),divide(b::'a,c))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1509
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1510
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1511
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1512
(*3160 inferences so far.  Searching to depth 11.  3.5 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1513
val HEN009_5 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1514
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1515
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1516
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1517
\  (\\<forall>Y X. equal(divide(divide(X::'a,Y),X),zero)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1518
\  (\\<forall>X Y Z. equal(divide(divide(divide(X::'a,Z),divide(Y::'a,Z)),divide(divide(X::'a,Y),Z)),zero)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1519
\  (\\<forall>X. equal(divide(zero::'a,X),zero)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1520
\  (\\<forall>X Y. equal(divide(X::'a,Y),zero) & equal(divide(Y::'a,X),zero) --> equal(X::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1521
\  (\\<forall>X. equal(divide(X::'a,identity),zero)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1522
\  (\\<forall>A B C. equal(A::'a,B) --> equal(divide(A::'a,C),divide(B::'a,C))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1523
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(divide(F'::'a,D),divide(F'::'a,E))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1524
\  (\\<forall>Y X Z. equal(divide(X::'a,Y),zero) & equal(divide(Y::'a,Z),zero) --> equal(divide(X::'a,Z),zero)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1525
\  (\\<forall>X Z Y. equal(divide(divide(X::'a,Y),Z),zero) --> equal(divide(divide(X::'a,Z),Y),zero)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1526
\  (\\<forall>Y Z X. equal(divide(X::'a,Y),zero) --> equal(divide(divide(Z::'a,Y),divide(Z::'a,X)),zero)) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1527
\  (~equal(divide(identity::'a,a),divide(identity::'a,divide(identity::'a,divide(identity::'a,a))))) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1528
\  (equal(divide(identity::'a,a),b)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1529
\  (equal(divide(identity::'a,b),c)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1530
\  (equal(divide(identity::'a,c),d)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1531
\  (~equal(b::'a,d)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1532
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1533
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1534
(*970373 inferences so far.  Searching to depth 17.  890.0 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1535
val HEN012_3 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1536
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1537
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1538
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1539
\  (\\<forall>X Y. less_equal(X::'a,Y) --> equal(divide(X::'a,Y),zero)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1540
\  (\\<forall>X Y. equal(divide(X::'a,Y),zero) --> less_equal(X::'a,Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1541
\  (\\<forall>Y X. less_equal(divide(X::'a,Y),X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1542
\  (\\<forall>X Y Z. less_equal(divide(divide(X::'a,Z),divide(Y::'a,Z)),divide(divide(X::'a,Y),Z))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1543
\  (\\<forall>X. less_equal(zero::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1544
\  (\\<forall>X Y. less_equal(X::'a,Y) & less_equal(Y::'a,X) --> equal(X::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1545
\  (\\<forall>X. less_equal(X::'a,identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1546
\  (\\<forall>A B C. equal(A::'a,B) --> equal(divide(A::'a,C),divide(B::'a,C))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1547
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(divide(F'::'a,D),divide(F'::'a,E))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1548
\  (\\<forall>G H I'. equal(G::'a,H) & less_equal(G::'a,I') --> less_equal(H::'a,I')) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1549
\  (\\<forall>J L K'. equal(J::'a,K') & less_equal(L::'a,J) --> less_equal(L::'a,K')) &     \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1550
\  (~less_equal(a::'a,a)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1551
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1552
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1553
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1554
(*1063 inferences so far.  Searching to depth 20.  1.0 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1555
val LCL010_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1556
 ("(\\<forall>X Y. is_a_theorem(equivalent(X::'a,Y)) & is_a_theorem(X) --> is_a_theorem(Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1557
\  (\\<forall>X Z Y. is_a_theorem(equivalent(equivalent(X::'a,Y),equivalent(equivalent(X::'a,Z),equivalent(Z::'a,Y))))) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1558
\  (~is_a_theorem(equivalent(equivalent(a::'a,b),equivalent(equivalent(c::'a,b),equivalent(a::'a,c))))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1559
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1560
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1561
(*2549 inferences so far.  Searching to depth 12.  1.4 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1562
val LCL077_2 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1563
 ("(\\<forall>X Y. is_a_theorem(implies(X,Y)) & is_a_theorem(X) --> is_a_theorem(Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1564
\  (\\<forall>Y X. is_a_theorem(implies(X,implies(Y,X)))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1565
\  (\\<forall>Y X Z. is_a_theorem(implies(implies(X,implies(Y,Z)),implies(implies(X,Y),implies(X,Z))))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1566
\  (\\<forall>Y X. is_a_theorem(implies(implies(not(X),not(Y)),implies(Y,X)))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1567
\  (\\<forall>X2 X1 X3. is_a_theorem(implies(X1,X2)) & is_a_theorem(implies(X2,X3)) --> is_a_theorem(implies(X1,X3))) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1568
\  (~is_a_theorem(implies(not(not(a)),a))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1569
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1570
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1571
(*2036 inferences so far.  Searching to depth 20.  1.5 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1572
val LCL082_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1573
 ("(\\<forall>X Y. is_a_theorem(implies(X::'a,Y)) & is_a_theorem(X) --> is_a_theorem(Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1574
\  (\\<forall>Y Z U X. is_a_theorem(implies(implies(implies(X::'a,Y),Z),implies(implies(Z::'a,X),implies(U::'a,X))))) &     \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1575
\  (~is_a_theorem(implies(a::'a,implies(b::'a,a)))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1576
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1577
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1578
(*1100 inferences so far.  Searching to depth 13.  1.0 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1579
val LCL111_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1580
 ("(\\<forall>X Y. is_a_theorem(implies(X,Y)) & is_a_theorem(X) --> is_a_theorem(Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1581
\  (\\<forall>Y X. is_a_theorem(implies(X,implies(Y,X)))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1582
\  (\\<forall>Y X Z. is_a_theorem(implies(implies(X,Y),implies(implies(Y,Z),implies(X,Z))))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1583
\  (\\<forall>Y X. is_a_theorem(implies(implies(implies(X,Y),Y),implies(implies(Y,X),X)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1584
\  (\\<forall>Y X. is_a_theorem(implies(implies(not(X),not(Y)),implies(Y,X)))) &        \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1585
\  (~is_a_theorem(implies(implies(a,b),implies(implies(c,a),implies(c,b))))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1586
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1587
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1588
(*667 inferences so far.  Searching to depth 9.  1.4 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1589
val LCL143_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1590
 ("(\\<forall>X. equal(X,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1591
\  (\\<forall>Y X. equal(X,Y) --> equal(Y,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1592
\  (\\<forall>Y X Z. equal(X,Y) & equal(Y,Z) --> equal(X,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1593
\  (\\<forall>X. equal(implies(true,X),X)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1594
\  (\\<forall>Y X Z. equal(implies(implies(X,Y),implies(implies(Y,Z),implies(X,Z))),true)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1595
\  (\\<forall>Y X. equal(implies(implies(X,Y),Y),implies(implies(Y,X),X))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1596
\  (\\<forall>Y X. equal(implies(implies(not(X),not(Y)),implies(Y,X)),true)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1597
\  (\\<forall>A B C. equal(A,B) --> equal(implies(A,C),implies(B,C))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1598
\  (\\<forall>D F' E. equal(D,E) --> equal(implies(F',D),implies(F',E))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1599
\  (\\<forall>G H. equal(G,H) --> equal(not(G),not(H))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1600
\  (\\<forall>X Y. equal(big_V(X,Y),implies(implies(X,Y),Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1601
\  (\\<forall>X Y. equal(big_hat(X,Y),not(big_V(not(X),not(Y))))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1602
\  (\\<forall>X Y. ordered(X,Y) --> equal(implies(X,Y),true)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1603
\  (\\<forall>X Y. equal(implies(X,Y),true) --> ordered(X,Y)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1604
\  (\\<forall>A B C. equal(A,B) --> equal(big_V(A,C),big_V(B,C))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1605
\  (\\<forall>D F' E. equal(D,E) --> equal(big_V(F',D),big_V(F',E))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1606
\  (\\<forall>G H I'. equal(G,H) --> equal(big_hat(G,I'),big_hat(H,I'))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1607
\  (\\<forall>J L K'. equal(J,K') --> equal(big_hat(L,J),big_hat(L,K'))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1608
\  (\\<forall>M N O_. equal(M,N) & ordered(M,O_) --> ordered(N,O_)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1609
\  (\\<forall>P R Q. equal(P,Q) & ordered(R,P) --> ordered(R,Q)) &      \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1610
\  (ordered(x,y)) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1611
\  (~ordered(implies(z,x),implies(z,y))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1612
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1613
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1614
(*5245 inferences so far.  Searching to depth 12.  4.6 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1615
val LCL182_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1616
 ("(\\<forall>A. axiom(or(not(or(A,A)),A))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1617
\  (\\<forall>B A. axiom(or(not(A),or(B,A)))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1618
\  (\\<forall>B A. axiom(or(not(or(A,B)),or(B,A)))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1619
\  (\\<forall>B A C. axiom(or(not(or(A,or(B,C))),or(B,or(A,C))))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1620
\  (\\<forall>A C B. axiom(or(not(or(not(A),B)),or(not(or(C,A)),or(C,B))))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1621
\  (\\<forall>X. axiom(X) --> theorem(X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1622
\  (\\<forall>X Y. axiom(or(not(Y),X)) & theorem(Y) --> theorem(X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1623
\  (\\<forall>X Y Z. axiom(or(not(X),Y)) & theorem(or(not(Y),Z)) --> theorem(or(not(X),Z))) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1624
\  (~theorem(or(not(or(not(p),q)),or(not(not(q)),not(p))))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1625
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1626
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1627
(*410 inferences so far.  Searching to depth 10.  0.3 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1628
val LCL200_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1629
 ("(\\<forall>A. axiom(or(not(or(A,A)),A))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1630
\  (\\<forall>B A. axiom(or(not(A),or(B,A)))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1631
\  (\\<forall>B A. axiom(or(not(or(A,B)),or(B,A)))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1632
\  (\\<forall>B A C. axiom(or(not(or(A,or(B,C))),or(B,or(A,C))))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1633
\  (\\<forall>A C B. axiom(or(not(or(not(A),B)),or(not(or(C,A)),or(C,B))))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1634
\  (\\<forall>X. axiom(X) --> theorem(X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1635
\  (\\<forall>X Y. axiom(or(not(Y),X)) & theorem(Y) --> theorem(X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1636
\  (\\<forall>X Y Z. axiom(or(not(X),Y)) & theorem(or(not(Y),Z)) --> theorem(or(not(X),Z))) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1637
\  (~theorem(or(not(not(or(p,q))),not(q)))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1638
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1639
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1640
(*5849 inferences so far.  Searching to depth 12.  5.6 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1641
val LCL215_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1642
 ("(\\<forall>A. axiom(or(not(or(A,A)),A))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1643
\  (\\<forall>B A. axiom(or(not(A),or(B,A)))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1644
\  (\\<forall>B A. axiom(or(not(or(A,B)),or(B,A)))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1645
\  (\\<forall>B A C. axiom(or(not(or(A,or(B,C))),or(B,or(A,C))))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1646
\  (\\<forall>A C B. axiom(or(not(or(not(A),B)),or(not(or(C,A)),or(C,B))))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1647
\  (\\<forall>X. axiom(X) --> theorem(X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1648
\  (\\<forall>X Y. axiom(or(not(Y),X)) & theorem(Y) --> theorem(X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1649
\  (\\<forall>X Y Z. axiom(or(not(X),Y)) & theorem(or(not(Y),Z)) --> theorem(or(not(X),Z))) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1650
\  (~theorem(or(not(or(not(p),q)),or(not(or(p,q)),q)))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1651
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1652
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1653
(*0 secs.  Not sure that a search even starts!*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1654
val LCL230_2 = prove
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1655
 ("(q --> p | r) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1656
\  (~p) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1657
\  (q) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1658
\  (~r) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1659
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1660
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1661
(*119585 inferences so far.  Searching to depth 14.  262.4 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1662
val LDA003_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1663
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1664
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1665
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1666
\  (\\<forall>Y X Z. equal(f(X::'a,f(Y::'a,Z)),f(f(X::'a,Y),f(X::'a,Z)))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1667
\  (\\<forall>X Y. left(X::'a,f(X::'a,Y))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1668
\  (\\<forall>Y X Z. left(X::'a,Y) & left(Y::'a,Z) --> left(X::'a,Z)) &     \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1669
\  (equal(num2::'a,f(num1::'a,num1))) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1670
\  (equal(num3::'a,f(num2::'a,num1))) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1671
\  (equal(u::'a,f(num2::'a,num2))) &  \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1672
\  (\\<forall>A B C. equal(A::'a,B) --> equal(f(A::'a,C),f(B::'a,C))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1673
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(f(F'::'a,D),f(F'::'a,E))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1674
\  (\\<forall>G H I'. equal(G::'a,H) & left(G::'a,I') --> left(H::'a,I')) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1675
\  (\\<forall>J L K'. equal(J::'a,K') & left(L::'a,J) --> left(L::'a,K')) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1676
\  (~left(num3::'a,u)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1677
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1678
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1679
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1680
(*2392 inferences so far.  Searching to depth 12.  2.2 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1681
val MSC002_1 = prove
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1682
 ("(at(something::'a,here,now)) &   \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1683
\  (\\<forall>Place Situation. hand_at(Place::'a,Situation) --> hand_at(Place::'a,let_go(Situation))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1684
\  (\\<forall>Place Another_place Situation. hand_at(Place::'a,Situation) --> hand_at(Another_place::'a,go(Another_place::'a,Situation))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1685
\  (\\<forall>Thing Situation. ~held(Thing::'a,let_go(Situation))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1686
\  (\\<forall>Situation Thing. at(Thing::'a,here,Situation) --> red(Thing)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1687
\  (\\<forall>Thing Place Situation. at(Thing::'a,Place,Situation) --> at(Thing::'a,Place,let_go(Situation))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1688
\  (\\<forall>Thing Place Situation. at(Thing::'a,Place,Situation) --> at(Thing::'a,Place,pick_up(Situation))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1689
\  (\\<forall>Thing Place Situation. at(Thing::'a,Place,Situation) --> grabbed(Thing::'a,pick_up(go(Place::'a,let_go(Situation))))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1690
\  (\\<forall>Thing Situation. red(Thing) & put(Thing::'a,there,Situation) --> answer(Situation)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1691
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1692
\  (\\<forall>Thing Place Another_place Situation. at(Thing::'a,Place,Situation) --> held(Thing::'a,Situation) | at(Thing::'a,Place,go(Another_place::'a,Situation))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1693
\  (\\<forall>One_place Thing Place Situation. hand_at(One_place::'a,Situation) & held(Thing::'a,Situation) --> at(Thing::'a,Place,go(Place::'a,Situation))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1694
\  (\\<forall>Place Thing Situation. hand_at(Place::'a,Situation) & at(Thing::'a,Place,Situation) --> held(Thing::'a,pick_up(Situation))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1695
\  (\\<forall>Situation. ~answer(Situation)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1696
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1697
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1698
(*73 inferences so far.  Searching to depth 10.  0.2 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1699
val MSC003_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1700
 ("(\\<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,times(Number_of_mid_parts::'a,Number_of_small_parts),Small_part)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1701
\  (\\<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,times(Number_of_mid_parts::'a,Number_of_small_parts),Small_part)) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1702
\  (in'(john::'a,boy)) &    \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1703
\  (\\<forall>X. in'(X::'a,boy) --> in'(X::'a,human)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1704
\  (\\<forall>X. in'(X::'a,hand) --> has_parts(X::'a,num5,fingers)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1705
\  (\\<forall>X. in'(X::'a,human) --> has_parts(X::'a,num2,arm)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1706
\  (\\<forall>X. in'(X::'a,arm) --> has_parts(X::'a,num1,hand)) &  \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1707
\  (~has_parts(john::'a,times(num2::'a,num1),hand)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1708
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1709
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1710
(*1486 inferences so far.  Searching to depth 20.  1.2 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1711
val MSC004_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1712
 ("(\\<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,times(Number_of_mid_parts::'a,Number_of_small_parts),Small_part)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1713
\  (\\<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,times(Number_of_mid_parts::'a,Number_of_small_parts),Small_part)) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1714
\  (in'(john::'a,boy)) &    \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1715
\  (\\<forall>X. in'(X::'a,boy) --> in'(X::'a,human)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1716
\  (\\<forall>X. in'(X::'a,hand) --> has_parts(X::'a,num5,fingers)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1717
\  (\\<forall>X. in'(X::'a,human) --> has_parts(X::'a,num2,arm)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1718
\  (\\<forall>X. in'(X::'a,arm) --> has_parts(X::'a,num1,hand)) &  \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1719
\  (~has_parts(john::'a,times(times(num2::'a,num1),num5),fingers)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1720
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1721
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1722
(*100 inferences so far.  Searching to depth 12.  0.1 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1723
val MSC005_1 = prove
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1724
 ("(value(truth::'a,truth)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1725
\  (value(falsity::'a,falsity)) &   \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1726
\  (\\<forall>X Y. value(X::'a,truth) & value(Y::'a,truth) --> value(xor(X::'a,Y),falsity)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1727
\  (\\<forall>X Y. value(X::'a,truth) & value(Y::'a,falsity) --> value(xor(X::'a,Y),truth)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1728
\  (\\<forall>X Y. value(X::'a,falsity) & value(Y::'a,truth) --> value(xor(X::'a,Y),truth)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1729
\  (\\<forall>X Y. value(X::'a,falsity) & value(Y::'a,falsity) --> value(xor(X::'a,Y),falsity)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1730
\  (\\<forall>Value. ~value(xor(xor(xor(xor(truth::'a,falsity),falsity),truth),falsity),Value)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1731
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1732
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1733
(*19116 inferences so far.  Searching to depth 16.  15.9 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1734
val MSC006_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1735
 ("(\\<forall>Y X Z. p(X::'a,Y) & p(Y::'a,Z) --> p(X::'a,Z)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1736
\  (\\<forall>Y X Z. q(X::'a,Y) & q(Y::'a,Z) --> q(X::'a,Z)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1737
\  (\\<forall>Y X. q(X::'a,Y) --> q(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1738
\  (\\<forall>X Y. p(X::'a,Y) | q(X::'a,Y)) &  \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1739
\  (~p(a::'a,b)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1740
\  (~q(c::'a,d)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1741
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1742
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1743
(*1713 inferences so far.  Searching to depth 10.  2.8 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1744
val NUM001_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1745
 ("(\\<forall>A. equal(A::'a,A)) &                                                \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1746
\  (\\<forall>B A C. equal(A::'a,B) & equal(B::'a,C) --> equal(A::'a,C)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1747
\  (\\<forall>B A. equal(add(A::'a,B),add(B::'a,A))) &                            \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1748
\  (\\<forall>A B C. equal(add(A::'a,add(B::'a,C)),add(add(A::'a,B),C))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1749
\  (\\<forall>B A. equal(subtract(add(A::'a,B),B),A)) &                           \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1750
\  (\\<forall>A B. equal(A::'a,subtract(add(A::'a,B),B))) &                       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1751
\  (\\<forall>A C B. equal(add(subtract(A::'a,B),C),subtract(add(A::'a,C),B))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1752
\  (\\<forall>A C B. equal(subtract(add(A::'a,B),C),add(subtract(A::'a,C),B))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1753
\  (\\<forall>A C B D. equal(A::'a,B) & equal(C::'a,add(A::'a,D)) --> equal(C::'a,add(B::'a,D))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1754
\  (\\<forall>A C D B. equal(A::'a,B) & equal(C::'a,add(D::'a,A)) --> equal(C::'a,add(D::'a,B))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1755
\  (\\<forall>A C B D. equal(A::'a,B) & equal(C::'a,subtract(A::'a,D)) --> equal(C::'a,subtract(B::'a,D))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1756
\  (\\<forall>A C D B. equal(A::'a,B) & equal(C::'a,subtract(D::'a,A)) --> equal(C::'a,subtract(D::'a,B))) &        \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1757
\  (~equal(add(add(a::'a,b),c),add(a::'a,add(b::'a,c)))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1758
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1759
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1760
(*20717 inferences so far.  Searching to depth 11.  13.7 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1761
val NUM021_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1762
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1763
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1764
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1765
\  (\\<forall>A. equal(add(A::'a,num0),A)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1766
\  (\\<forall>A B. equal(add(A::'a,successor(B)),successor(add(A::'a,B)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1767
\  (\\<forall>A. equal(multiply(A::'a,num0),num0)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1768
\  (\\<forall>B A. equal(multiply(A::'a,successor(B)),add(multiply(A::'a,B),A))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1769
\  (\\<forall>A B. equal(successor(A),successor(B)) --> equal(A::'a,B)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1770
\  (\\<forall>A B. equal(A::'a,B) --> equal(successor(A),successor(B))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1771
\  (\\<forall>A C B. less(A::'a,B) & less(C::'a,A) --> less(C::'a,B)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1772
\  (\\<forall>A B C. equal(add(successor(A),B),C) --> less(B::'a,C)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1773
\  (\\<forall>A B. less(A::'a,B) --> equal(add(successor(predecessor_of_1st_minus_2nd(B::'a,A)),A),B)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1774
\  (\\<forall>A B. divides(A::'a,B) --> less(A::'a,B) | equal(A::'a,B)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1775
\  (\\<forall>A B. less(A::'a,B) --> divides(A::'a,B)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1776
\  (\\<forall>A B. equal(A::'a,B) --> divides(A::'a,B)) &       \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1777
\  (less(b::'a,c)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1778
\  (~less(b::'a,a)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1779
\  (divides(c::'a,a)) &     \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1780
\  (\\<forall>A. ~equal(successor(A),num0)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1781
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1782
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1783
(*26320 inferences so far.  Searching to depth 10.  26.4 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1784
val NUM024_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1785
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1786
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1787
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1788
\  (\\<forall>A. equal(add(A::'a,num0),A)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1789
\  (\\<forall>A B. equal(add(A::'a,successor(B)),successor(add(A::'a,B)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1790
\  (\\<forall>A. equal(multiply(A::'a,num0),num0)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1791
\  (\\<forall>B A. equal(multiply(A::'a,successor(B)),add(multiply(A::'a,B),A))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1792
\  (\\<forall>A B. equal(successor(A),successor(B)) --> equal(A::'a,B)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1793
\  (\\<forall>A B. equal(A::'a,B) --> equal(successor(A),successor(B))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1794
\  (\\<forall>A C B. less(A::'a,B) & less(C::'a,A) --> less(C::'a,B)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1795
\  (\\<forall>A B C. equal(add(successor(A),B),C) --> less(B::'a,C)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1796
\  (\\<forall>A B. less(A::'a,B) --> equal(add(successor(predecessor_of_1st_minus_2nd(B::'a,A)),A),B)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1797
\  (\\<forall>B A. equal(add(A::'a,B),add(B::'a,A))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1798
\  (\\<forall>B A C. equal(add(A::'a,B),add(C::'a,B)) --> equal(A::'a,C)) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1799
\  (less(a::'a,a)) &        \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1800
\  (\\<forall>A. ~equal(successor(A),num0)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  1801
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1802
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1803
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1804
(*1345 inferences so far.  Searching to depth 7.  23.3 secs.  BIG*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1805
val NUM180_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1806
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1807
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1808
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1809
\  (\\<forall>X U Y. subclass(X::'a,Y) & member(U::'a,X) --> member(U::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1810
\  (\\<forall>X Y. member(not_subclass_element(X::'a,Y),X) | subclass(X::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1811
\  (\\<forall>X Y. member(not_subclass_element(X::'a,Y),Y) --> subclass(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1812
\  (\\<forall>X. subclass(X::'a,universal_class)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1813
\  (\\<forall>X Y. equal(X::'a,Y) --> subclass(X::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1814
\  (\\<forall>Y X. equal(X::'a,Y) --> subclass(Y::'a,X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1815
\  (\\<forall>X Y. subclass(X::'a,Y) & subclass(Y::'a,X) --> equal(X::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1816
\  (\\<forall>X U Y. member(U::'a,unordered_pair(X::'a,Y)) --> equal(U::'a,X) | equal(U::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1817
\  (\\<forall>X Y. member(X::'a,universal_class) --> member(X::'a,unordered_pair(X::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1818
\  (\\<forall>X Y. member(Y::'a,universal_class) --> member(Y::'a,unordered_pair(X::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1819
\  (\\<forall>X Y. member(unordered_pair(X::'a,Y),universal_class)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1820
\  (\\<forall>X. equal(unordered_pair(X::'a,X),singleton(X))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1821
\  (\\<forall>X Y. equal(unordered_pair(singleton(X),unordered_pair(X::'a,singleton(Y))),ordered_pair(X::'a,Y))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1822
\  (\\<forall>V Y U X. member(ordered_pair(U::'a,V),cross_product(X::'a,Y)) --> member(U::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1823
\  (\\<forall>U X V Y. member(ordered_pair(U::'a,V),cross_product(X::'a,Y)) --> member(V::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1824
\  (\\<forall>U V X Y. member(U::'a,X) & member(V::'a,Y) --> member(ordered_pair(U::'a,V),cross_product(X::'a,Y))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1825
\  (\\<forall>X Y Z. member(Z::'a,cross_product(X::'a,Y)) --> equal(ordered_pair(first(Z),second(Z)),Z)) &      \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1826
\  (subclass(element_relation::'a,cross_product(universal_class::'a,universal_class))) &        \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1827
\  (\\<forall>X Y. member(ordered_pair(X::'a,Y),element_relation) --> member(X::'a,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1828
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1829
\  (\\<forall>Y Z X. member(Z::'a,intersection(X::'a,Y)) --> member(Z::'a,X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1830
\  (\\<forall>X Z Y. member(Z::'a,intersection(X::'a,Y)) --> member(Z::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1831
\  (\\<forall>Z X Y. member(Z::'a,X) & member(Z::'a,Y) --> member(Z::'a,intersection(X::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1832
\  (\\<forall>Z X. ~(member(Z::'a,complement(X)) & member(Z::'a,X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1833
\  (\\<forall>Z X. member(Z::'a,universal_class) --> member(Z::'a,complement(X)) | member(Z::'a,X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1834
\  (\\<forall>X Y. equal(complement(intersection(complement(X),complement(Y))),union(X::'a,Y))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1835
\  (\\<forall>X Y. equal(intersection(complement(intersection(X::'a,Y)),complement(intersection(complement(X),complement(Y)))),difference(X::'a,Y))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1836
\  (\\<forall>Xr X Y. equal(intersection(Xr::'a,cross_product(X::'a,Y)),restrct(Xr::'a,X,Y))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1837
\  (\\<forall>Xr X Y. equal(intersection(cross_product(X::'a,Y),Xr),restrct(Xr::'a,X,Y))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1838
\  (\\<forall>Z X. ~(equal(restrct(X::'a,singleton(Z),universal_class),null_class) & member(Z::'a,domain_of(X)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1839
\  (\\<forall>Z X. member(Z::'a,universal_class) --> equal(restrct(X::'a,singleton(Z),universal_class),null_class) | member(Z::'a,domain_of(X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1840
\  (\\<forall>X. subclass(rotate(X),cross_product(cross_product(universal_class::'a,universal_class),universal_class))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1841
\  (\\<forall>V W U X. member(ordered_pair(ordered_pair(U::'a,V),W),rotate(X)) --> member(ordered_pair(ordered_pair(V::'a,W),U),X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1842
\  (\\<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),rotate(X))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1843
\  (\\<forall>X. subclass(flip(X),cross_product(cross_product(universal_class::'a,universal_class),universal_class))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1844
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1845
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1846
\  (\\<forall>Y. equal(domain_of(flip(cross_product(Y::'a,universal_class))),INVERSE(Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1847
\  (\\<forall>Z. equal(domain_of(INVERSE(Z)),range_of(Z))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1848
\  (\\<forall>Z X Y. equal(first(not_subclass_element(restrct(Z::'a,X,singleton(Y)),null_class)),domain(Z::'a,X,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1849
\  (\\<forall>Z X Y. equal(second(not_subclass_element(restrct(Z::'a,singleton(X),Y),null_class)),rng(Z::'a,X,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1850
\  (\\<forall>Xr X. equal(range_of(restrct(Xr::'a,X,universal_class)),image_(Xr::'a,X))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1851
\  (\\<forall>X. equal(union(X::'a,singleton(X)),successor(X))) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1852
\  (subclass(successor_relation::'a,cross_product(universal_class::'a,universal_class))) &      \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1853
\  (\\<forall>X Y. member(ordered_pair(X::'a,Y),successor_relation) --> equal(successor(X),Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1854
\  (\\<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)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1855
\  (\\<forall>X. inductive(X) --> member(null_class::'a,X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1856
\  (\\<forall>X. inductive(X) --> subclass(image_(successor_relation::'a,X),X)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1857
\  (\\<forall>X. member(null_class::'a,X) & subclass(image_(successor_relation::'a,X),X) --> inductive(X)) &     \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1858
\  (inductive(omega)) & \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1859
\  (\\<forall>Y. inductive(Y) --> subclass(omega::'a,Y)) &  \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1860
\  (member(omega::'a,universal_class)) &    \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1861
\  (\\<forall>X. equal(domain_of(restrct(element_relation::'a,universal_class,X)),sum_class(X))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1862
\  (\\<forall>X. member(X::'a,universal_class) --> member(sum_class(X),universal_class)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1863
\  (\\<forall>X. equal(complement(image_(element_relation::'a,complement(X))),powerClass(X))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1864
\  (\\<forall>U. member(U::'a,universal_class) --> member(powerClass(U),universal_class)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1865
\  (\\<forall>Yr Xr. subclass(compos(Yr::'a,Xr),cross_product(universal_class::'a,universal_class))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1866
\  (\\<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))))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1867
\  (\\<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))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1868
\  (\\<forall>X. single_valued_class(X) --> subclass(compos(X::'a,INVERSE(X)),identity_relation)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1869
\  (\\<forall>X. subclass(compos(X::'a,INVERSE(X)),identity_relation) --> single_valued_class(X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1870
\  (\\<forall>Xf. function(Xf) --> subclass(Xf::'a,cross_product(universal_class::'a,universal_class))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1871
\  (\\<forall>Xf. function(Xf) --> subclass(compos(Xf::'a,INVERSE(Xf)),identity_relation)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1872
\  (\\<forall>Xf. subclass(Xf::'a,cross_product(universal_class::'a,universal_class)) & subclass(compos(Xf::'a,INVERSE(Xf)),identity_relation) --> function(Xf)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1873
\  (\\<forall>Xf X. function(Xf) & member(X::'a,universal_class) --> member(image_(Xf::'a,X),universal_class)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1874
\  (\\<forall>X. equal(X::'a,null_class) | member(regular(X),X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1875
\  (\\<forall>X. equal(X::'a,null_class) | equal(intersection(X::'a,regular(X)),null_class)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1876
\  (\\<forall>Xf Y. equal(sum_class(image_(Xf::'a,singleton(Y))),apply(Xf::'a,Y))) &     \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1877
\  (function(choice)) & \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1878
\  (\\<forall>Y. member(Y::'a,universal_class) --> equal(Y::'a,null_class) | member(apply(choice::'a,Y),Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1879
\  (\\<forall>Xf. one_to_one(Xf) --> function(Xf)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1880
\  (\\<forall>Xf. one_to_one(Xf) --> function(INVERSE(Xf))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1881
\  (\\<forall>Xf. function(INVERSE(Xf)) & function(Xf) --> one_to_one(Xf)) &    \
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  1882
\  (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)) &     \
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  1883
\  (equal(intersection(INVERSE(subset_relation),subset_relation),identity_relation)) &  \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1884
\  (\\<forall>Xr. equal(complement(domain_of(intersection(Xr::'a,identity_relation))),diagonalise(Xr))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1885
\  (\\<forall>X. equal(intersection(domain_of(X),diagonalise(compos(INVERSE(element_relation),X))),cantor(X))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1886
\  (\\<forall>Xf. operation(Xf) --> function(Xf)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1887
\  (\\<forall>Xf. operation(Xf) --> equal(cross_product(domain_of(domain_of(Xf)),domain_of(domain_of(Xf))),domain_of(Xf))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1888
\  (\\<forall>Xf. operation(Xf) --> subclass(range_of(Xf),domain_of(domain_of(Xf)))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1889
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1890
\  (\\<forall>Xf1 Xf2 Xh. compatible(Xh::'a,Xf1,Xf2) --> function(Xh)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1891
\  (\\<forall>Xf2 Xf1 Xh. compatible(Xh::'a,Xf1,Xf2) --> equal(domain_of(domain_of(Xf1)),domain_of(Xh))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1892
\  (\\<forall>Xf1 Xh Xf2. compatible(Xh::'a,Xf1,Xf2) --> subclass(range_of(Xh),domain_of(domain_of(Xf2)))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1893
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1894
\  (\\<forall>Xh Xf2 Xf1. homomorphism(Xh::'a,Xf1,Xf2) --> operation(Xf1)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1895
\  (\\<forall>Xh Xf1 Xf2. homomorphism(Xh::'a,Xf1,Xf2) --> operation(Xf2)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1896
\  (\\<forall>Xh Xf1 Xf2. homomorphism(Xh::'a,Xf1,Xf2) --> compatible(Xh::'a,Xf1,Xf2)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1897
\  (\\<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))))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1898
\  (\\<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)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1899
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1900
\  (\\<forall>D E F'. equal(D::'a,E) --> equal(apply(D::'a,F'),apply(E::'a,F'))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1901
\  (\\<forall>G I' H. equal(G::'a,H) --> equal(apply(I'::'a,G),apply(I'::'a,H))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1902
\  (\\<forall>J K'. equal(J::'a,K') --> equal(cantor(J),cantor(K'))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1903
\  (\\<forall>L M. equal(L::'a,M) --> equal(complement(L),complement(M))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1904
\  (\\<forall>N O_ P. equal(N::'a,O_) --> equal(compos(N::'a,P),compos(O_::'a,P))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1905
\  (\\<forall>Q S' R. equal(Q::'a,R) --> equal(compos(S'::'a,Q),compos(S'::'a,R))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1906
\  (\\<forall>T' U V. equal(T'::'a,U) --> equal(cross_product(T'::'a,V),cross_product(U::'a,V))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1907
\  (\\<forall>W Y X. equal(W::'a,X) --> equal(cross_product(Y::'a,W),cross_product(Y::'a,X))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1908
\  (\\<forall>Z A1. equal(Z::'a,A1) --> equal(diagonalise(Z),diagonalise(A1))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1909
\  (\\<forall>B1 C1 D1. equal(B1::'a,C1) --> equal(difference(B1::'a,D1),difference(C1::'a,D1))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1910
\  (\\<forall>E1 G1 F1. equal(E1::'a,F1) --> equal(difference(G1::'a,E1),difference(G1::'a,F1))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1911
\  (\\<forall>H1 I1 J1 K1. equal(H1::'a,I1) --> equal(domain(H1::'a,J1,K1),domain(I1::'a,J1,K1))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1912
\  (\\<forall>L1 N1 M1 O1. equal(L1::'a,M1) --> equal(domain(N1::'a,L1,O1),domain(N1::'a,M1,O1))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1913
\  (\\<forall>P1 R1 S1 Q1. equal(P1::'a,Q1) --> equal(domain(R1::'a,S1,P1),domain(R1::'a,S1,Q1))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1914
\  (\\<forall>T1 U1. equal(T1::'a,U1) --> equal(domain_of(T1),domain_of(U1))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1915
\  (\\<forall>V1 W1. equal(V1::'a,W1) --> equal(first(V1),first(W1))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1916
\  (\\<forall>X1 Y1. equal(X1::'a,Y1) --> equal(flip(X1),flip(Y1))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1917
\  (\\<forall>Z1 A2 B2. equal(Z1::'a,A2) --> equal(image_(Z1::'a,B2),image_(A2::'a,B2))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1918
\  (\\<forall>C2 E2 D2. equal(C2::'a,D2) --> equal(image_(E2::'a,C2),image_(E2::'a,D2))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1919
\  (\\<forall>F2 G2 H2. equal(F2::'a,G2) --> equal(intersection(F2::'a,H2),intersection(G2::'a,H2))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1920
\  (\\<forall>I2 K2 J2. equal(I2::'a,J2) --> equal(intersection(K2::'a,I2),intersection(K2::'a,J2))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1921
\  (\\<forall>L2 M2. equal(L2::'a,M2) --> equal(INVERSE(L2),INVERSE(M2))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1922
\  (\\<forall>N2 O2 P2 Q2. equal(N2::'a,O2) --> equal(not_homomorphism1(N2::'a,P2,Q2),not_homomorphism1(O2::'a,P2,Q2))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1923
\  (\\<forall>R2 T2 S2 U2. equal(R2::'a,S2) --> equal(not_homomorphism1(T2::'a,R2,U2),not_homomorphism1(T2::'a,S2,U2))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1924
\  (\\<forall>V2 X2 Y2 W2. equal(V2::'a,W2) --> equal(not_homomorphism1(X2::'a,Y2,V2),not_homomorphism1(X2::'a,Y2,W2))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1925
\  (\\<forall>Z2 A3 B3 C3. equal(Z2::'a,A3) --> equal(not_homomorphism2(Z2::'a,B3,C3),not_homomorphism2(A3::'a,B3,C3))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1926
\  (\\<forall>D3 F3 E3 G3. equal(D3::'a,E3) --> equal(not_homomorphism2(F3::'a,D3,G3),not_homomorphism2(F3::'a,E3,G3))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1927
\  (\\<forall>H3 J3 K3 I3. equal(H3::'a,I3) --> equal(not_homomorphism2(J3::'a,K3,H3),not_homomorphism2(J3::'a,K3,I3))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1928
\  (\\<forall>L3 M3 N3. equal(L3::'a,M3) --> equal(not_subclass_element(L3::'a,N3),not_subclass_element(M3::'a,N3))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1929
\  (\\<forall>O3 Q3 P3. equal(O3::'a,P3) --> equal(not_subclass_element(Q3::'a,O3),not_subclass_element(Q3::'a,P3))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1930
\  (\\<forall>R3 S3 T3. equal(R3::'a,S3) --> equal(ordered_pair(R3::'a,T3),ordered_pair(S3::'a,T3))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1931
\  (\\<forall>U3 W3 V3. equal(U3::'a,V3) --> equal(ordered_pair(W3::'a,U3),ordered_pair(W3::'a,V3))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1932
\  (\\<forall>X3 Y3. equal(X3::'a,Y3) --> equal(powerClass(X3),powerClass(Y3))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1933
\  (\\<forall>Z3 A4 B4 C4. equal(Z3::'a,A4) --> equal(rng(Z3::'a,B4,C4),rng(A4::'a,B4,C4))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1934
\  (\\<forall>D4 F4 E4 G4. equal(D4::'a,E4) --> equal(rng(F4::'a,D4,G4),rng(F4::'a,E4,G4))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1935
\  (\\<forall>H4 J4 K4 I4. equal(H4::'a,I4) --> equal(rng(J4::'a,K4,H4),rng(J4::'a,K4,I4))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1936
\  (\\<forall>L4 M4. equal(L4::'a,M4) --> equal(range_of(L4),range_of(M4))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1937
\  (\\<forall>N4 O4. equal(N4::'a,O4) --> equal(regular(N4),regular(O4))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1938
\  (\\<forall>P4 Q4 R4 S4. equal(P4::'a,Q4) --> equal(restrct(P4::'a,R4,S4),restrct(Q4::'a,R4,S4))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1939
\  (\\<forall>T4 V4 U4 W4. equal(T4::'a,U4) --> equal(restrct(V4::'a,T4,W4),restrct(V4::'a,U4,W4))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1940
\  (\\<forall>X4 Z4 A5 Y4. equal(X4::'a,Y4) --> equal(restrct(Z4::'a,A5,X4),restrct(Z4::'a,A5,Y4))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1941
\  (\\<forall>B5 C5. equal(B5::'a,C5) --> equal(rotate(B5),rotate(C5))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1942
\  (\\<forall>D5 E5. equal(D5::'a,E5) --> equal(second(D5),second(E5))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1943
\  (\\<forall>F5 G5. equal(F5::'a,G5) --> equal(singleton(F5),singleton(G5))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1944
\  (\\<forall>H5 I5. equal(H5::'a,I5) --> equal(successor(H5),successor(I5))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1945
\  (\\<forall>J5 K5. equal(J5::'a,K5) --> equal(sum_class(J5),sum_class(K5))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1946
\  (\\<forall>L5 M5 N5. equal(L5::'a,M5) --> equal(union(L5::'a,N5),union(M5::'a,N5))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1947
\  (\\<forall>O5 Q5 P5. equal(O5::'a,P5) --> equal(union(Q5::'a,O5),union(Q5::'a,P5))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1948
\  (\\<forall>R5 S5 T5. equal(R5::'a,S5) --> equal(unordered_pair(R5::'a,T5),unordered_pair(S5::'a,T5))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1949
\  (\\<forall>U5 W5 V5. equal(U5::'a,V5) --> equal(unordered_pair(W5::'a,U5),unordered_pair(W5::'a,V5))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1950
\  (\\<forall>X5 Y5 Z5 A6. equal(X5::'a,Y5) & compatible(X5::'a,Z5,A6) --> compatible(Y5::'a,Z5,A6)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1951
\  (\\<forall>B6 D6 C6 E6. equal(B6::'a,C6) & compatible(D6::'a,B6,E6) --> compatible(D6::'a,C6,E6)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1952
\  (\\<forall>F6 H6 I6 G6. equal(F6::'a,G6) & compatible(H6::'a,I6,F6) --> compatible(H6::'a,I6,G6)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1953
\  (\\<forall>J6 K6. equal(J6::'a,K6) & function(J6) --> function(K6)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1954
\  (\\<forall>L6 M6 N6 O6. equal(L6::'a,M6) & homomorphism(L6::'a,N6,O6) --> homomorphism(M6::'a,N6,O6)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1955
\  (\\<forall>P6 R6 Q6 S6. equal(P6::'a,Q6) & homomorphism(R6::'a,P6,S6) --> homomorphism(R6::'a,Q6,S6)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1956
\  (\\<forall>T6 V6 W6 U6. equal(T6::'a,U6) & homomorphism(V6::'a,W6,T6) --> homomorphism(V6::'a,W6,U6)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1957
\  (\\<forall>X6 Y6. equal(X6::'a,Y6) & inductive(X6) --> inductive(Y6)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1958
\  (\\<forall>Z6 A7 B7. equal(Z6::'a,A7) & member(Z6::'a,B7) --> member(A7::'a,B7)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1959
\  (\\<forall>C7 E7 D7. equal(C7::'a,D7) & member(E7::'a,C7) --> member(E7::'a,D7)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1960
\  (\\<forall>F7 G7. equal(F7::'a,G7) & one_to_one(F7) --> one_to_one(G7)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1961
\  (\\<forall>H7 I7. equal(H7::'a,I7) & operation(H7) --> operation(I7)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1962
\  (\\<forall>J7 K7. equal(J7::'a,K7) & single_valued_class(J7) --> single_valued_class(K7)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1963
\  (\\<forall>L7 M7 N7. equal(L7::'a,M7) & subclass(L7::'a,N7) --> subclass(M7::'a,N7)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1964
\  (\\<forall>O7 Q7 P7. equal(O7::'a,P7) & subclass(Q7::'a,O7) --> subclass(Q7::'a,P7)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1965
\  (\\<forall>X. subclass(compose_class(X),cross_product(universal_class::'a,universal_class))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1966
\  (\\<forall>X Y Z. member(ordered_pair(Y::'a,Z),compose_class(X)) --> equal(compos(X::'a,Y),Z)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1967
\  (\\<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))) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1968
\  (subclass(composition_function::'a,cross_product(universal_class::'a,cross_product(universal_class::'a,universal_class)))) &     \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1969
\  (\\<forall>X Y Z. member(ordered_pair(X::'a,ordered_pair(Y::'a,Z)),composition_function) --> equal(compos(X::'a,Y),Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1970
\  (\\<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)) &    \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1971
\  (subclass(domain_relation::'a,cross_product(universal_class::'a,universal_class))) & \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1972
\  (\\<forall>X Y. member(ordered_pair(X::'a,Y),domain_relation) --> equal(domain_of(X),Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1973
\  (\\<forall>X. member(X::'a,universal_class) --> member(ordered_pair(X::'a,domain_of(X)),domain_relation)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1974
\  (\\<forall>X. equal(first(not_subclass_element(compos(X::'a,INVERSE(X)),identity_relation)),single_valued1(X))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1975
\  (\\<forall>X. equal(second(not_subclass_element(compos(X::'a,INVERSE(X)),identity_relation)),single_valued2(X))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1976
\  (\\<forall>X. equal(domain(X::'a,image_(INVERSE(X),singleton(single_valued1(X))),single_valued2(X)),single_valued3(X))) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1977
\  (equal(intersection(complement(compos(element_relation::'a,complement(identity_relation))),element_relation),singleton_relation)) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  1978
\  (subclass(application_function::'a,cross_product(universal_class::'a,cross_product(universal_class::'a,universal_class)))) &     \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1979
\  (\\<forall>Z Y X. member(ordered_pair(X::'a,ordered_pair(Y::'a,Z)),application_function) --> member(Y::'a,domain_of(X))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1980
\  (\\<forall>X Y Z. member(ordered_pair(X::'a,ordered_pair(Y::'a,Z)),application_function) --> equal(apply(X::'a,Y),Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1981
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1982
\  (\\<forall>X Y Xf. maps(Xf::'a,X,Y) --> function(Xf)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1983
\  (\\<forall>Y Xf X. maps(Xf::'a,X,Y) --> equal(domain_of(Xf),X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1984
\  (\\<forall>X Xf Y. maps(Xf::'a,X,Y) --> subclass(range_of(Xf),Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1985
\  (\\<forall>Xf Y. function(Xf) & subclass(range_of(Xf),Y) --> maps(Xf::'a,domain_of(Xf),Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1986
\  (\\<forall>L M. equal(L::'a,M) --> equal(compose_class(L),compose_class(M))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1987
\  (\\<forall>N2 O2. equal(N2::'a,O2) --> equal(single_valued1(N2),single_valued1(O2))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1988
\  (\\<forall>P2 Q2. equal(P2::'a,Q2) --> equal(single_valued2(P2),single_valued2(Q2))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1989
\  (\\<forall>R2 S2. equal(R2::'a,S2) --> equal(single_valued3(R2),single_valued3(S2))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1990
\  (\\<forall>X2 Y2 Z2 A3. equal(X2::'a,Y2) & maps(X2::'a,Z2,A3) --> maps(Y2::'a,Z2,A3)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1991
\  (\\<forall>B3 D3 C3 E3. equal(B3::'a,C3) & maps(D3::'a,B3,E3) --> maps(D3::'a,C3,E3)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1992
\  (\\<forall>F3 H3 I3 G3. equal(F3::'a,G3) & maps(H3::'a,I3,F3) --> maps(H3::'a,I3,G3)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1993
\  (\\<forall>X. equal(union(X::'a,INVERSE(X)),symmetrization_of(X))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1994
\  (\\<forall>X Y. irreflexive(X::'a,Y) --> subclass(restrct(X::'a,Y,Y),complement(identity_relation))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1995
\  (\\<forall>X Y. subclass(restrct(X::'a,Y,Y),complement(identity_relation)) --> irreflexive(X::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1996
\  (\\<forall>Y X. connected(X::'a,Y) --> subclass(cross_product(Y::'a,Y),union(identity_relation::'a,symmetrization_of(X)))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1997
\  (\\<forall>X Y. subclass(cross_product(Y::'a,Y),union(identity_relation::'a,symmetrization_of(X))) --> connected(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1998
\  (\\<forall>Xr Y. transitive(Xr::'a,Y) --> subclass(compos(restrct(Xr::'a,Y,Y),restrct(Xr::'a,Y,Y)),restrct(Xr::'a,Y,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  1999
\  (\\<forall>Xr Y. subclass(compos(restrct(Xr::'a,Y,Y),restrct(Xr::'a,Y,Y)),restrct(Xr::'a,Y,Y)) --> transitive(Xr::'a,Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2000
\  (\\<forall>Xr Y. asymmetric(Xr::'a,Y) --> equal(restrct(intersection(Xr::'a,INVERSE(Xr)),Y,Y),null_class)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2001
\  (\\<forall>Xr Y. equal(restrct(intersection(Xr::'a,INVERSE(Xr)),Y,Y),null_class) --> asymmetric(Xr::'a,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2002
\  (\\<forall>Xr Y Z. equal(segment(Xr::'a,Y,Z),domain_of(restrct(Xr::'a,Y,singleton(Z))))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2003
\  (\\<forall>X Y. well_ordering(X::'a,Y) --> connected(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2004
\  (\\<forall>Y Xr U. well_ordering(Xr::'a,Y) & subclass(U::'a,Y) --> equal(U::'a,null_class) | member(least(Xr::'a,U),U)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2005
\  (\\<forall>Y V Xr U. well_ordering(Xr::'a,Y) & subclass(U::'a,Y) & member(V::'a,U) --> member(least(Xr::'a,U),U)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2006
\  (\\<forall>Y Xr U. well_ordering(Xr::'a,Y) & subclass(U::'a,Y) --> equal(segment(Xr::'a,U,least(Xr::'a,U)),null_class)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2007
\  (\\<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))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2008
\  (\\<forall>Xr Y. connected(Xr::'a,Y) & equal(not_well_ordering(Xr::'a,Y),null_class) --> well_ordering(Xr::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2009
\  (\\<forall>Xr Y. connected(Xr::'a,Y) --> subclass(not_well_ordering(Xr::'a,Y),Y) | well_ordering(Xr::'a,Y)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2010
\  (\\<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)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2011
\  (\\<forall>Xr Y Z. section(Xr::'a,Y,Z) --> subclass(Y::'a,Z)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2012
\  (\\<forall>Xr Z Y. section(Xr::'a,Y,Z) --> subclass(domain_of(restrct(Xr::'a,Z,Y)),Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2013
\  (\\<forall>Xr Y Z. subclass(Y::'a,Z) & subclass(domain_of(restrct(Xr::'a,Z,Y)),Y) --> section(Xr::'a,Y,Z)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2014
\  (\\<forall>X. member(X::'a,ordinal_numbers) --> well_ordering(element_relation::'a,X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2015
\  (\\<forall>X. member(X::'a,ordinal_numbers) --> subclass(sum_class(X),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2016
\  (\\<forall>X. well_ordering(element_relation::'a,X) & subclass(sum_class(X),X) & member(X::'a,universal_class) --> member(X::'a,ordinal_numbers)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2017
\  (\\<forall>X. well_ordering(element_relation::'a,X) & subclass(sum_class(X),X) --> member(X::'a,ordinal_numbers) | equal(X::'a,ordinal_numbers)) &      \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2018
\  (equal(union(singleton(null_class),image_(successor_relation::'a,ordinal_numbers)),kind_1_ordinals)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2019
\  (equal(intersection(complement(kind_1_ordinals),ordinal_numbers),limit_ordinals)) &  \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2020
\  (\\<forall>X. subclass(rest_of(X),cross_product(universal_class::'a,universal_class))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2021
\  (\\<forall>V U X. member(ordered_pair(U::'a,V),rest_of(X)) --> member(U::'a,domain_of(X))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2022
\  (\\<forall>X U V. member(ordered_pair(U::'a,V),rest_of(X)) --> equal(restrct(X::'a,U,universal_class),V)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2023
\  (\\<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))) &        \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2024
\  (subclass(rest_relation::'a,cross_product(universal_class::'a,universal_class))) &   \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2025
\  (\\<forall>X Y. member(ordered_pair(X::'a,Y),rest_relation) --> equal(rest_of(X),Y)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2026
\  (\\<forall>X. member(X::'a,universal_class) --> member(ordered_pair(X::'a,rest_of(X)),rest_relation)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2027
\  (\\<forall>X Z. member(X::'a,recursion_equation_functions(Z)) --> function(Z)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2028
\  (\\<forall>Z X. member(X::'a,recursion_equation_functions(Z)) --> function(X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2029
\  (\\<forall>Z X. member(X::'a,recursion_equation_functions(Z)) --> member(domain_of(X),ordinal_numbers)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2030
\  (\\<forall>Z X. member(X::'a,recursion_equation_functions(Z)) --> equal(compos(Z::'a,rest_of(X)),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2031
\  (\\<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))) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2032
\  (subclass(union_of_range_map::'a,cross_product(universal_class::'a,universal_class))) &      \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2033
\  (\\<forall>X Y. member(ordered_pair(X::'a,Y),union_of_range_map) --> equal(sum_class(range_of(X)),Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2034
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2035
\  (\\<forall>X Y. equal(apply(recursion(X::'a,successor_relation,union_of_range_map),Y),ordinal_add(X::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2036
\  (\\<forall>X Y. equal(recursion(null_class::'a,apply(add_relation::'a,X),union_of_range_map),ordinal_multiply(X::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2037
\  (\\<forall>X. member(X::'a,omega) --> equal(integer_of(X),X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2038
\  (\\<forall>X. member(X::'a,omega) | equal(integer_of(X),null_class)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2039
\  (\\<forall>D E. equal(D::'a,E) --> equal(integer_of(D),integer_of(E))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2040
\  (\\<forall>F' G H. equal(F'::'a,G) --> equal(least(F'::'a,H),least(G::'a,H))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2041
\  (\\<forall>I' K' J. equal(I'::'a,J) --> equal(least(K'::'a,I'),least(K'::'a,J))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2042
\  (\\<forall>L M N. equal(L::'a,M) --> equal(not_well_ordering(L::'a,N),not_well_ordering(M::'a,N))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2043
\  (\\<forall>O_ Q P. equal(O_::'a,P) --> equal(not_well_ordering(Q::'a,O_),not_well_ordering(Q::'a,P))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2044
\  (\\<forall>R S' T'. equal(R::'a,S') --> equal(ordinal_add(R::'a,T'),ordinal_add(S'::'a,T'))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2045
\  (\\<forall>U W V. equal(U::'a,V) --> equal(ordinal_add(W::'a,U),ordinal_add(W::'a,V))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2046
\  (\\<forall>X Y Z. equal(X::'a,Y) --> equal(ordinal_multiply(X::'a,Z),ordinal_multiply(Y::'a,Z))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2047
\  (\\<forall>A1 C1 B1. equal(A1::'a,B1) --> equal(ordinal_multiply(C1::'a,A1),ordinal_multiply(C1::'a,B1))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2048
\  (\\<forall>F1 G1 H1 I1. equal(F1::'a,G1) --> equal(recursion(F1::'a,H1,I1),recursion(G1::'a,H1,I1))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2049
\  (\\<forall>J1 L1 K1 M1. equal(J1::'a,K1) --> equal(recursion(L1::'a,J1,M1),recursion(L1::'a,K1,M1))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2050
\  (\\<forall>N1 P1 Q1 O1. equal(N1::'a,O1) --> equal(recursion(P1::'a,Q1,N1),recursion(P1::'a,Q1,O1))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2051
\  (\\<forall>R1 S1. equal(R1::'a,S1) --> equal(recursion_equation_functions(R1),recursion_equation_functions(S1))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2052
\  (\\<forall>T1 U1. equal(T1::'a,U1) --> equal(rest_of(T1),rest_of(U1))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2053
\  (\\<forall>V1 W1 X1 Y1. equal(V1::'a,W1) --> equal(segment(V1::'a,X1,Y1),segment(W1::'a,X1,Y1))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2054
\  (\\<forall>Z1 B2 A2 C2. equal(Z1::'a,A2) --> equal(segment(B2::'a,Z1,C2),segment(B2::'a,A2,C2))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2055
\  (\\<forall>D2 F2 G2 E2. equal(D2::'a,E2) --> equal(segment(F2::'a,G2,D2),segment(F2::'a,G2,E2))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2056
\  (\\<forall>H2 I2. equal(H2::'a,I2) --> equal(symmetrization_of(H2),symmetrization_of(I2))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2057
\  (\\<forall>J2 K2 L2. equal(J2::'a,K2) & asymmetric(J2::'a,L2) --> asymmetric(K2::'a,L2)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2058
\  (\\<forall>M2 O2 N2. equal(M2::'a,N2) & asymmetric(O2::'a,M2) --> asymmetric(O2::'a,N2)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2059
\  (\\<forall>P2 Q2 R2. equal(P2::'a,Q2) & connected(P2::'a,R2) --> connected(Q2::'a,R2)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2060
\  (\\<forall>S2 U2 T2. equal(S2::'a,T2) & connected(U2::'a,S2) --> connected(U2::'a,T2)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2061
\  (\\<forall>V2 W2 X2. equal(V2::'a,W2) & irreflexive(V2::'a,X2) --> irreflexive(W2::'a,X2)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2062
\  (\\<forall>Y2 A3 Z2. equal(Y2::'a,Z2) & irreflexive(A3::'a,Y2) --> irreflexive(A3::'a,Z2)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2063
\  (\\<forall>B3 C3 D3 E3. equal(B3::'a,C3) & section(B3::'a,D3,E3) --> section(C3::'a,D3,E3)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2064
\  (\\<forall>F3 H3 G3 I3. equal(F3::'a,G3) & section(H3::'a,F3,I3) --> section(H3::'a,G3,I3)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2065
\  (\\<forall>J3 L3 M3 K3. equal(J3::'a,K3) & section(L3::'a,M3,J3) --> section(L3::'a,M3,K3)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2066
\  (\\<forall>N3 O3 P3. equal(N3::'a,O3) & transitive(N3::'a,P3) --> transitive(O3::'a,P3)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2067
\  (\\<forall>Q3 S3 R3. equal(Q3::'a,R3) & transitive(S3::'a,Q3) --> transitive(S3::'a,R3)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2068
\  (\\<forall>T3 U3 V3. equal(T3::'a,U3) & well_ordering(T3::'a,V3) --> well_ordering(U3::'a,V3)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2069
\  (\\<forall>W3 Y3 X3. equal(W3::'a,X3) & well_ordering(Y3::'a,W3) --> well_ordering(Y3::'a,X3)) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2070
\  (~subclass(limit_ordinals::'a,ordinal_numbers)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2071
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2072
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2073
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2074
(*0 inferences so far.  Searching to depth 0.  16.8 secs.  BIG*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2075
val NUM228_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2076
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2077
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2078
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2079
\  (\\<forall>X U Y. subclass(X::'a,Y) & member(U::'a,X) --> member(U::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2080
\  (\\<forall>X Y. member(not_subclass_element(X::'a,Y),X) | subclass(X::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2081
\  (\\<forall>X Y. member(not_subclass_element(X::'a,Y),Y) --> subclass(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2082
\  (\\<forall>X. subclass(X::'a,universal_class)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2083
\  (\\<forall>X Y. equal(X::'a,Y) --> subclass(X::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2084
\  (\\<forall>Y X. equal(X::'a,Y) --> subclass(Y::'a,X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2085
\  (\\<forall>X Y. subclass(X::'a,Y) & subclass(Y::'a,X) --> equal(X::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2086
\  (\\<forall>X U Y. member(U::'a,unordered_pair(X::'a,Y)) --> equal(U::'a,X) | equal(U::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2087
\  (\\<forall>X Y. member(X::'a,universal_class) --> member(X::'a,unordered_pair(X::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2088
\  (\\<forall>X Y. member(Y::'a,universal_class) --> member(Y::'a,unordered_pair(X::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2089
\  (\\<forall>X Y. member(unordered_pair(X::'a,Y),universal_class)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2090
\  (\\<forall>X. equal(unordered_pair(X::'a,X),singleton(X))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2091
\  (\\<forall>X Y. equal(unordered_pair(singleton(X),unordered_pair(X::'a,singleton(Y))),ordered_pair(X::'a,Y))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2092
\  (\\<forall>V Y U X. member(ordered_pair(U::'a,V),cross_product(X::'a,Y)) --> member(U::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2093
\  (\\<forall>U X V Y. member(ordered_pair(U::'a,V),cross_product(X::'a,Y)) --> member(V::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2094
\  (\\<forall>U V X Y. member(U::'a,X) & member(V::'a,Y) --> member(ordered_pair(U::'a,V),cross_product(X::'a,Y))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2095
\  (\\<forall>X Y Z. member(Z::'a,cross_product(X::'a,Y)) --> equal(ordered_pair(first(Z),second(Z)),Z)) &      \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2096
\  (subclass(element_relation::'a,cross_product(universal_class::'a,universal_class))) &        \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2097
\  (\\<forall>X Y. member(ordered_pair(X::'a,Y),element_relation) --> member(X::'a,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2098
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2099
\  (\\<forall>Y Z X. member(Z::'a,intersection(X::'a,Y)) --> member(Z::'a,X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2100
\  (\\<forall>X Z Y. member(Z::'a,intersection(X::'a,Y)) --> member(Z::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2101
\  (\\<forall>Z X Y. member(Z::'a,X) & member(Z::'a,Y) --> member(Z::'a,intersection(X::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2102
\  (\\<forall>Z X. ~(member(Z::'a,complement(X)) & member(Z::'a,X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2103
\  (\\<forall>Z X. member(Z::'a,universal_class) --> member(Z::'a,complement(X)) | member(Z::'a,X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2104
\  (\\<forall>X Y. equal(complement(intersection(complement(X),complement(Y))),union(X::'a,Y))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2105
\  (\\<forall>X Y. equal(intersection(complement(intersection(X::'a,Y)),complement(intersection(complement(X),complement(Y)))),difference(X::'a,Y))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2106
\  (\\<forall>Xr X Y. equal(intersection(Xr::'a,cross_product(X::'a,Y)),restrct(Xr::'a,X,Y))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2107
\  (\\<forall>Xr X Y. equal(intersection(cross_product(X::'a,Y),Xr),restrct(Xr::'a,X,Y))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2108
\  (\\<forall>Z X. ~(equal(restrct(X::'a,singleton(Z),universal_class),null_class) & member(Z::'a,domain_of(X)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2109
\  (\\<forall>Z X. member(Z::'a,universal_class) --> equal(restrct(X::'a,singleton(Z),universal_class),null_class) | member(Z::'a,domain_of(X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2110
\  (\\<forall>X. subclass(rotate(X),cross_product(cross_product(universal_class::'a,universal_class),universal_class))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2111
\  (\\<forall>V W U X. member(ordered_pair(ordered_pair(U::'a,V),W),rotate(X)) --> member(ordered_pair(ordered_pair(V::'a,W),U),X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2112
\  (\\<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),rotate(X))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2113
\  (\\<forall>X. subclass(flip(X),cross_product(cross_product(universal_class::'a,universal_class),universal_class))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2114
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2115
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2116
\  (\\<forall>Y. equal(domain_of(flip(cross_product(Y::'a,universal_class))),INVERSE(Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2117
\  (\\<forall>Z. equal(domain_of(INVERSE(Z)),range_of(Z))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2118
\  (\\<forall>Z X Y. equal(first(not_subclass_element(restrct(Z::'a,X,singleton(Y)),null_class)),domain(Z::'a,X,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2119
\  (\\<forall>Z X Y. equal(second(not_subclass_element(restrct(Z::'a,singleton(X),Y),null_class)),rng(Z::'a,X,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2120
\  (\\<forall>Xr X. equal(range_of(restrct(Xr::'a,X,universal_class)),image_(Xr::'a,X))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2121
\  (\\<forall>X. equal(union(X::'a,singleton(X)),successor(X))) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2122
\  (subclass(successor_relation::'a,cross_product(universal_class::'a,universal_class))) &      \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2123
\  (\\<forall>X Y. member(ordered_pair(X::'a,Y),successor_relation) --> equal(successor(X),Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2124
\  (\\<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)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2125
\  (\\<forall>X. inductive(X) --> member(null_class::'a,X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2126
\  (\\<forall>X. inductive(X) --> subclass(image_(successor_relation::'a,X),X)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2127
\  (\\<forall>X. member(null_class::'a,X) & subclass(image_(successor_relation::'a,X),X) --> inductive(X)) &     \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2128
\  (inductive(omega)) & \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2129
\  (\\<forall>Y. inductive(Y) --> subclass(omega::'a,Y)) &  \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2130
\  (member(omega::'a,universal_class)) &    \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2131
\  (\\<forall>X. equal(domain_of(restrct(element_relation::'a,universal_class,X)),sum_class(X))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2132
\  (\\<forall>X. member(X::'a,universal_class) --> member(sum_class(X),universal_class)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2133
\  (\\<forall>X. equal(complement(image_(element_relation::'a,complement(X))),powerClass(X))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2134
\  (\\<forall>U. member(U::'a,universal_class) --> member(powerClass(U),universal_class)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2135
\  (\\<forall>Yr Xr. subclass(compos(Yr::'a,Xr),cross_product(universal_class::'a,universal_class))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2136
\  (\\<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))))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2137
\  (\\<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))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2138
\  (\\<forall>X. single_valued_class(X) --> subclass(compos(X::'a,INVERSE(X)),identity_relation)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2139
\  (\\<forall>X. subclass(compos(X::'a,INVERSE(X)),identity_relation) --> single_valued_class(X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2140
\  (\\<forall>Xf. function(Xf) --> subclass(Xf::'a,cross_product(universal_class::'a,universal_class))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2141
\  (\\<forall>Xf. function(Xf) --> subclass(compos(Xf::'a,INVERSE(Xf)),identity_relation)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2142
\  (\\<forall>Xf. subclass(Xf::'a,cross_product(universal_class::'a,universal_class)) & subclass(compos(Xf::'a,INVERSE(Xf)),identity_relation) --> function(Xf)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2143
\  (\\<forall>Xf X. function(Xf) & member(X::'a,universal_class) --> member(image_(Xf::'a,X),universal_class)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2144
\  (\\<forall>X. equal(X::'a,null_class) | member(regular(X),X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2145
\  (\\<forall>X. equal(X::'a,null_class) | equal(intersection(X::'a,regular(X)),null_class)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2146
\  (\\<forall>Xf Y. equal(sum_class(image_(Xf::'a,singleton(Y))),apply(Xf::'a,Y))) &     \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2147
\  (function(choice)) & \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2148
\  (\\<forall>Y. member(Y::'a,universal_class) --> equal(Y::'a,null_class) | member(apply(choice::'a,Y),Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2149
\  (\\<forall>Xf. one_to_one(Xf) --> function(Xf)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2150
\  (\\<forall>Xf. one_to_one(Xf) --> function(INVERSE(Xf))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2151
\  (\\<forall>Xf. function(INVERSE(Xf)) & function(Xf) --> one_to_one(Xf)) &    \
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2152
\  (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)) &     \
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2153
\  (equal(intersection(INVERSE(subset_relation),subset_relation),identity_relation)) &  \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2154
\  (\\<forall>Xr. equal(complement(domain_of(intersection(Xr::'a,identity_relation))),diagonalise(Xr))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2155
\  (\\<forall>X. equal(intersection(domain_of(X),diagonalise(compos(INVERSE(element_relation),X))),cantor(X))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2156
\  (\\<forall>Xf. operation(Xf) --> function(Xf)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2157
\  (\\<forall>Xf. operation(Xf) --> equal(cross_product(domain_of(domain_of(Xf)),domain_of(domain_of(Xf))),domain_of(Xf))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2158
\  (\\<forall>Xf. operation(Xf) --> subclass(range_of(Xf),domain_of(domain_of(Xf)))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2159
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2160
\  (\\<forall>Xf1 Xf2 Xh. compatible(Xh::'a,Xf1,Xf2) --> function(Xh)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2161
\  (\\<forall>Xf2 Xf1 Xh. compatible(Xh::'a,Xf1,Xf2) --> equal(domain_of(domain_of(Xf1)),domain_of(Xh))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2162
\  (\\<forall>Xf1 Xh Xf2. compatible(Xh::'a,Xf1,Xf2) --> subclass(range_of(Xh),domain_of(domain_of(Xf2)))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2163
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2164
\  (\\<forall>Xh Xf2 Xf1. homomorphism(Xh::'a,Xf1,Xf2) --> operation(Xf1)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2165
\  (\\<forall>Xh Xf1 Xf2. homomorphism(Xh::'a,Xf1,Xf2) --> operation(Xf2)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2166
\  (\\<forall>Xh Xf1 Xf2. homomorphism(Xh::'a,Xf1,Xf2) --> compatible(Xh::'a,Xf1,Xf2)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2167
\  (\\<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))))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2168
\  (\\<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)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2169
\  (\\<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)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2170
\  (\\<forall>D E F'. equal(D::'a,E) --> equal(apply(D::'a,F'),apply(E::'a,F'))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2171
\  (\\<forall>G I' H. equal(G::'a,H) --> equal(apply(I'::'a,G),apply(I'::'a,H))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2172
\  (\\<forall>J K'. equal(J::'a,K') --> equal(cantor(J),cantor(K'))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2173
\  (\\<forall>L M. equal(L::'a,M) --> equal(complement(L),complement(M))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2174
\  (\\<forall>N O_ P. equal(N::'a,O_) --> equal(compos(N::'a,P),compos(O_::'a,P))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2175
\  (\\<forall>Q S' R. equal(Q::'a,R) --> equal(compos(S'::'a,Q),compos(S'::'a,R))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2176
\  (\\<forall>T' U V. equal(T'::'a,U) --> equal(cross_product(T'::'a,V),cross_product(U::'a,V))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2177
\  (\\<forall>W Y X. equal(W::'a,X) --> equal(cross_product(Y::'a,W),cross_product(Y::'a,X))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2178
\  (\\<forall>Z A1. equal(Z::'a,A1) --> equal(diagonalise(Z),diagonalise(A1))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2179
\  (\\<forall>B1 C1 D1. equal(B1::'a,C1) --> equal(difference(B1::'a,D1),difference(C1::'a,D1))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2180
\  (\\<forall>E1 G1 F1. equal(E1::'a,F1) --> equal(difference(G1::'a,E1),difference(G1::'a,F1))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2181
\  (\\<forall>H1 I1 J1 K1. equal(H1::'a,I1) --> equal(domain(H1::'a,J1,K1),domain(I1::'a,J1,K1))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2182
\  (\\<forall>L1 N1 M1 O1. equal(L1::'a,M1) --> equal(domain(N1::'a,L1,O1),domain(N1::'a,M1,O1))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2183
\  (\\<forall>P1 R1 S1 Q1. equal(P1::'a,Q1) --> equal(domain(R1::'a,S1,P1),domain(R1::'a,S1,Q1))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2184
\  (\\<forall>T1 U1. equal(T1::'a,U1) --> equal(domain_of(T1),domain_of(U1))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2185
\  (\\<forall>V1 W1. equal(V1::'a,W1) --> equal(first(V1),first(W1))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2186
\  (\\<forall>X1 Y1. equal(X1::'a,Y1) --> equal(flip(X1),flip(Y1))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2187
\  (\\<forall>Z1 A2 B2. equal(Z1::'a,A2) --> equal(image_(Z1::'a,B2),image_(A2::'a,B2))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2188
\  (\\<forall>C2 E2 D2. equal(C2::'a,D2) --> equal(image_(E2::'a,C2),image_(E2::'a,D2))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2189
\  (\\<forall>F2 G2 H2. equal(F2::'a,G2) --> equal(intersection(F2::'a,H2),intersection(G2::'a,H2))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2190
\  (\\<forall>I2 K2 J2. equal(I2::'a,J2) --> equal(intersection(K2::'a,I2),intersection(K2::'a,J2))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2191
\  (\\<forall>L2 M2. equal(L2::'a,M2) --> equal(INVERSE(L2),INVERSE(M2))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2192
\  (\\<forall>N2 O2 P2 Q2. equal(N2::'a,O2) --> equal(not_homomorphism1(N2::'a,P2,Q2),not_homomorphism1(O2::'a,P2,Q2))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2193
\  (\\<forall>R2 T2 S2 U2. equal(R2::'a,S2) --> equal(not_homomorphism1(T2::'a,R2,U2),not_homomorphism1(T2::'a,S2,U2))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2194
\  (\\<forall>V2 X2 Y2 W2. equal(V2::'a,W2) --> equal(not_homomorphism1(X2::'a,Y2,V2),not_homomorphism1(X2::'a,Y2,W2))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2195
\  (\\<forall>Z2 A3 B3 C3. equal(Z2::'a,A3) --> equal(not_homomorphism2(Z2::'a,B3,C3),not_homomorphism2(A3::'a,B3,C3))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2196
\  (\\<forall>D3 F3 E3 G3. equal(D3::'a,E3) --> equal(not_homomorphism2(F3::'a,D3,G3),not_homomorphism2(F3::'a,E3,G3))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2197
\  (\\<forall>H3 J3 K3 I3. equal(H3::'a,I3) --> equal(not_homomorphism2(J3::'a,K3,H3),not_homomorphism2(J3::'a,K3,I3))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2198
\  (\\<forall>L3 M3 N3. equal(L3::'a,M3) --> equal(not_subclass_element(L3::'a,N3),not_subclass_element(M3::'a,N3))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2199
\  (\\<forall>O3 Q3 P3. equal(O3::'a,P3) --> equal(not_subclass_element(Q3::'a,O3),not_subclass_element(Q3::'a,P3))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2200
\  (\\<forall>R3 S3 T3. equal(R3::'a,S3) --> equal(ordered_pair(R3::'a,T3),ordered_pair(S3::'a,T3))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2201
\  (\\<forall>U3 W3 V3. equal(U3::'a,V3) --> equal(ordered_pair(W3::'a,U3),ordered_pair(W3::'a,V3))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2202
\  (\\<forall>X3 Y3. equal(X3::'a,Y3) --> equal(powerClass(X3),powerClass(Y3))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2203
\  (\\<forall>Z3 A4 B4 C4. equal(Z3::'a,A4) --> equal(rng(Z3::'a,B4,C4),rng(A4::'a,B4,C4))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2204
\  (\\<forall>D4 F4 E4 G4. equal(D4::'a,E4) --> equal(rng(F4::'a,D4,G4),rng(F4::'a,E4,G4))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2205
\  (\\<forall>H4 J4 K4 I4. equal(H4::'a,I4) --> equal(rng(J4::'a,K4,H4),rng(J4::'a,K4,I4))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2206
\  (\\<forall>L4 M4. equal(L4::'a,M4) --> equal(range_of(L4),range_of(M4))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2207
\  (\\<forall>N4 O4. equal(N4::'a,O4) --> equal(regular(N4),regular(O4))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2208
\  (\\<forall>P4 Q4 R4 S4. equal(P4::'a,Q4) --> equal(restrct(P4::'a,R4,S4),restrct(Q4::'a,R4,S4))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2209
\  (\\<forall>T4 V4 U4 W4. equal(T4::'a,U4) --> equal(restrct(V4::'a,T4,W4),restrct(V4::'a,U4,W4))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2210
\  (\\<forall>X4 Z4 A5 Y4. equal(X4::'a,Y4) --> equal(restrct(Z4::'a,A5,X4),restrct(Z4::'a,A5,Y4))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2211
\  (\\<forall>B5 C5. equal(B5::'a,C5) --> equal(rotate(B5),rotate(C5))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2212
\  (\\<forall>D5 E5. equal(D5::'a,E5) --> equal(second(D5),second(E5))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2213
\  (\\<forall>F5 G5. equal(F5::'a,G5) --> equal(singleton(F5),singleton(G5))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2214
\  (\\<forall>H5 I5. equal(H5::'a,I5) --> equal(successor(H5),successor(I5))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2215
\  (\\<forall>J5 K5. equal(J5::'a,K5) --> equal(sum_class(J5),sum_class(K5))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2216
\  (\\<forall>L5 M5 N5. equal(L5::'a,M5) --> equal(union(L5::'a,N5),union(M5::'a,N5))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2217
\  (\\<forall>O5 Q5 P5. equal(O5::'a,P5) --> equal(union(Q5::'a,O5),union(Q5::'a,P5))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2218
\  (\\<forall>R5 S5 T5. equal(R5::'a,S5) --> equal(unordered_pair(R5::'a,T5),unordered_pair(S5::'a,T5))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2219
\  (\\<forall>U5 W5 V5. equal(U5::'a,V5) --> equal(unordered_pair(W5::'a,U5),unordered_pair(W5::'a,V5))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2220
\  (\\<forall>X5 Y5 Z5 A6. equal(X5::'a,Y5) & compatible(X5::'a,Z5,A6) --> compatible(Y5::'a,Z5,A6)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2221
\  (\\<forall>B6 D6 C6 E6. equal(B6::'a,C6) & compatible(D6::'a,B6,E6) --> compatible(D6::'a,C6,E6)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2222
\  (\\<forall>F6 H6 I6 G6. equal(F6::'a,G6) & compatible(H6::'a,I6,F6) --> compatible(H6::'a,I6,G6)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2223
\  (\\<forall>J6 K6. equal(J6::'a,K6) & function(J6) --> function(K6)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2224
\  (\\<forall>L6 M6 N6 O6. equal(L6::'a,M6) & homomorphism(L6::'a,N6,O6) --> homomorphism(M6::'a,N6,O6)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2225
\  (\\<forall>P6 R6 Q6 S6. equal(P6::'a,Q6) & homomorphism(R6::'a,P6,S6) --> homomorphism(R6::'a,Q6,S6)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2226
\  (\\<forall>T6 V6 W6 U6. equal(T6::'a,U6) & homomorphism(V6::'a,W6,T6) --> homomorphism(V6::'a,W6,U6)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2227
\  (\\<forall>X6 Y6. equal(X6::'a,Y6) & inductive(X6) --> inductive(Y6)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2228
\  (\\<forall>Z6 A7 B7. equal(Z6::'a,A7) & member(Z6::'a,B7) --> member(A7::'a,B7)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2229
\  (\\<forall>C7 E7 D7. equal(C7::'a,D7) & member(E7::'a,C7) --> member(E7::'a,D7)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2230
\  (\\<forall>F7 G7. equal(F7::'a,G7) & one_to_one(F7) --> one_to_one(G7)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2231
\  (\\<forall>H7 I7. equal(H7::'a,I7) & operation(H7) --> operation(I7)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2232
\  (\\<forall>J7 K7. equal(J7::'a,K7) & single_valued_class(J7) --> single_valued_class(K7)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2233
\  (\\<forall>L7 M7 N7. equal(L7::'a,M7) & subclass(L7::'a,N7) --> subclass(M7::'a,N7)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2234
\  (\\<forall>O7 Q7 P7. equal(O7::'a,P7) & subclass(Q7::'a,O7) --> subclass(Q7::'a,P7)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2235
\  (\\<forall>X. subclass(compose_class(X),cross_product(universal_class::'a,universal_class))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2236
\  (\\<forall>X Y Z. member(ordered_pair(Y::'a,Z),compose_class(X)) --> equal(compos(X::'a,Y),Z)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2237
\  (\\<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))) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2238
\  (subclass(composition_function::'a,cross_product(universal_class::'a,cross_product(universal_class::'a,universal_class)))) &     \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2239
\  (\\<forall>X Y Z. member(ordered_pair(X::'a,ordered_pair(Y::'a,Z)),composition_function) --> equal(compos(X::'a,Y),Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2240
\  (\\<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)) &    \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2241
\  (subclass(domain_relation::'a,cross_product(universal_class::'a,universal_class))) & \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2242
\  (\\<forall>X Y. member(ordered_pair(X::'a,Y),domain_relation) --> equal(domain_of(X),Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2243
\  (\\<forall>X. member(X::'a,universal_class) --> member(ordered_pair(X::'a,domain_of(X)),domain_relation)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2244
\  (\\<forall>X. equal(first(not_subclass_element(compos(X::'a,INVERSE(X)),identity_relation)),single_valued1(X))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2245
\  (\\<forall>X. equal(second(not_subclass_element(compos(X::'a,INVERSE(X)),identity_relation)),single_valued2(X))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2246
\  (\\<forall>X. equal(domain(X::'a,image_(INVERSE(X),singleton(single_valued1(X))),single_valued2(X)),single_valued3(X))) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2247
\  (equal(intersection(complement(compos(element_relation::'a,complement(identity_relation))),element_relation),singleton_relation)) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2248
\  (subclass(application_function::'a,cross_product(universal_class::'a,cross_product(universal_class::'a,universal_class)))) &     \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2249
\  (\\<forall>Z Y X. member(ordered_pair(X::'a,ordered_pair(Y::'a,Z)),application_function) --> member(Y::'a,domain_of(X))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2250
\  (\\<forall>X Y Z. member(ordered_pair(X::'a,ordered_pair(Y::'a,Z)),application_function) --> equal(apply(X::'a,Y),Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2251
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2252
\  (\\<forall>X Y Xf. maps(Xf::'a,X,Y) --> function(Xf)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2253
\  (\\<forall>Y Xf X. maps(Xf::'a,X,Y) --> equal(domain_of(Xf),X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2254
\  (\\<forall>X Xf Y. maps(Xf::'a,X,Y) --> subclass(range_of(Xf),Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2255
\  (\\<forall>Xf Y. function(Xf) & subclass(range_of(Xf),Y) --> maps(Xf::'a,domain_of(Xf),Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2256
\  (\\<forall>L M. equal(L::'a,M) --> equal(compose_class(L),compose_class(M))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2257
\  (\\<forall>N2 O2. equal(N2::'a,O2) --> equal(single_valued1(N2),single_valued1(O2))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2258
\  (\\<forall>P2 Q2. equal(P2::'a,Q2) --> equal(single_valued2(P2),single_valued2(Q2))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2259
\  (\\<forall>R2 S2. equal(R2::'a,S2) --> equal(single_valued3(R2),single_valued3(S2))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2260
\  (\\<forall>X2 Y2 Z2 A3. equal(X2::'a,Y2) & maps(X2::'a,Z2,A3) --> maps(Y2::'a,Z2,A3)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2261
\  (\\<forall>B3 D3 C3 E3. equal(B3::'a,C3) & maps(D3::'a,B3,E3) --> maps(D3::'a,C3,E3)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2262
\  (\\<forall>F3 H3 I3 G3. equal(F3::'a,G3) & maps(H3::'a,I3,F3) --> maps(H3::'a,I3,G3)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2263
\  (\\<forall>X. equal(union(X::'a,INVERSE(X)),symmetrization_of(X))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2264
\  (\\<forall>X Y. irreflexive(X::'a,Y) --> subclass(restrct(X::'a,Y,Y),complement(identity_relation))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2265
\  (\\<forall>X Y. subclass(restrct(X::'a,Y,Y),complement(identity_relation)) --> irreflexive(X::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2266
\  (\\<forall>Y X. connected(X::'a,Y) --> subclass(cross_product(Y::'a,Y),union(identity_relation::'a,symmetrization_of(X)))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2267
\  (\\<forall>X Y. subclass(cross_product(Y::'a,Y),union(identity_relation::'a,symmetrization_of(X))) --> connected(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2268
\  (\\<forall>Xr Y. transitive(Xr::'a,Y) --> subclass(compos(restrct(Xr::'a,Y,Y),restrct(Xr::'a,Y,Y)),restrct(Xr::'a,Y,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2269
\  (\\<forall>Xr Y. subclass(compos(restrct(Xr::'a,Y,Y),restrct(Xr::'a,Y,Y)),restrct(Xr::'a,Y,Y)) --> transitive(Xr::'a,Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2270
\  (\\<forall>Xr Y. asymmetric(Xr::'a,Y) --> equal(restrct(intersection(Xr::'a,INVERSE(Xr)),Y,Y),null_class)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2271
\  (\\<forall>Xr Y. equal(restrct(intersection(Xr::'a,INVERSE(Xr)),Y,Y),null_class) --> asymmetric(Xr::'a,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2272
\  (\\<forall>Xr Y Z. equal(segment(Xr::'a,Y,Z),domain_of(restrct(Xr::'a,Y,singleton(Z))))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2273
\  (\\<forall>X Y. well_ordering(X::'a,Y) --> connected(X::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2274
\  (\\<forall>Y Xr U. well_ordering(Xr::'a,Y) & subclass(U::'a,Y) --> equal(U::'a,null_class) | member(least(Xr::'a,U),U)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2275
\  (\\<forall>Y V Xr U. well_ordering(Xr::'a,Y) & subclass(U::'a,Y) & member(V::'a,U) --> member(least(Xr::'a,U),U)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2276
\  (\\<forall>Y Xr U. well_ordering(Xr::'a,Y) & subclass(U::'a,Y) --> equal(segment(Xr::'a,U,least(Xr::'a,U)),null_class)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2277
\  (\\<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))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2278
\  (\\<forall>Xr Y. connected(Xr::'a,Y) & equal(not_well_ordering(Xr::'a,Y),null_class) --> well_ordering(Xr::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2279
\  (\\<forall>Xr Y. connected(Xr::'a,Y) --> subclass(not_well_ordering(Xr::'a,Y),Y) | well_ordering(Xr::'a,Y)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2280
\  (\\<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)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2281
\  (\\<forall>Xr Y Z. section(Xr::'a,Y,Z) --> subclass(Y::'a,Z)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2282
\  (\\<forall>Xr Z Y. section(Xr::'a,Y,Z) --> subclass(domain_of(restrct(Xr::'a,Z,Y)),Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2283
\  (\\<forall>Xr Y Z. subclass(Y::'a,Z) & subclass(domain_of(restrct(Xr::'a,Z,Y)),Y) --> section(Xr::'a,Y,Z)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2284
\  (\\<forall>X. member(X::'a,ordinal_numbers) --> well_ordering(element_relation::'a,X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2285
\  (\\<forall>X. member(X::'a,ordinal_numbers) --> subclass(sum_class(X),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2286
\  (\\<forall>X. well_ordering(element_relation::'a,X) & subclass(sum_class(X),X) & member(X::'a,universal_class) --> member(X::'a,ordinal_numbers)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2287
\  (\\<forall>X. well_ordering(element_relation::'a,X) & subclass(sum_class(X),X) --> member(X::'a,ordinal_numbers) | equal(X::'a,ordinal_numbers)) &      \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2288
\  (equal(union(singleton(null_class),image_(successor_relation::'a,ordinal_numbers)),kind_1_ordinals)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2289
\  (equal(intersection(complement(kind_1_ordinals),ordinal_numbers),limit_ordinals)) &  \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2290
\  (\\<forall>X. subclass(rest_of(X),cross_product(universal_class::'a,universal_class))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2291
\  (\\<forall>V U X. member(ordered_pair(U::'a,V),rest_of(X)) --> member(U::'a,domain_of(X))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2292
\  (\\<forall>X U V. member(ordered_pair(U::'a,V),rest_of(X)) --> equal(restrct(X::'a,U,universal_class),V)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2293
\  (\\<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))) &        \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2294
\  (subclass(rest_relation::'a,cross_product(universal_class::'a,universal_class))) &   \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2295
\  (\\<forall>X Y. member(ordered_pair(X::'a,Y),rest_relation) --> equal(rest_of(X),Y)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2296
\  (\\<forall>X. member(X::'a,universal_class) --> member(ordered_pair(X::'a,rest_of(X)),rest_relation)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2297
\  (\\<forall>X Z. member(X::'a,recursion_equation_functions(Z)) --> function(Z)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2298
\  (\\<forall>Z X. member(X::'a,recursion_equation_functions(Z)) --> function(X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2299
\  (\\<forall>Z X. member(X::'a,recursion_equation_functions(Z)) --> member(domain_of(X),ordinal_numbers)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2300
\  (\\<forall>Z X. member(X::'a,recursion_equation_functions(Z)) --> equal(compos(Z::'a,rest_of(X)),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2301
\  (\\<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))) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2302
\  (subclass(union_of_range_map::'a,cross_product(universal_class::'a,universal_class))) &      \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2303
\  (\\<forall>X Y. member(ordered_pair(X::'a,Y),union_of_range_map) --> equal(sum_class(range_of(X)),Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2304
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2305
\  (\\<forall>X Y. equal(apply(recursion(X::'a,successor_relation,union_of_range_map),Y),ordinal_add(X::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2306
\  (\\<forall>X Y. equal(recursion(null_class::'a,apply(add_relation::'a,X),union_of_range_map),ordinal_multiply(X::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2307
\  (\\<forall>X. member(X::'a,omega) --> equal(integer_of(X),X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2308
\  (\\<forall>X. member(X::'a,omega) | equal(integer_of(X),null_class)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2309
\  (\\<forall>D E. equal(D::'a,E) --> equal(integer_of(D),integer_of(E))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2310
\  (\\<forall>F' G H. equal(F'::'a,G) --> equal(least(F'::'a,H),least(G::'a,H))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2311
\  (\\<forall>I' K' J. equal(I'::'a,J) --> equal(least(K'::'a,I'),least(K'::'a,J))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2312
\  (\\<forall>L M N. equal(L::'a,M) --> equal(not_well_ordering(L::'a,N),not_well_ordering(M::'a,N))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2313
\  (\\<forall>O_ Q P. equal(O_::'a,P) --> equal(not_well_ordering(Q::'a,O_),not_well_ordering(Q::'a,P))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2314
\  (\\<forall>R S' T'. equal(R::'a,S') --> equal(ordinal_add(R::'a,T'),ordinal_add(S'::'a,T'))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2315
\  (\\<forall>U W V. equal(U::'a,V) --> equal(ordinal_add(W::'a,U),ordinal_add(W::'a,V))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2316
\  (\\<forall>X Y Z. equal(X::'a,Y) --> equal(ordinal_multiply(X::'a,Z),ordinal_multiply(Y::'a,Z))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2317
\  (\\<forall>A1 C1 B1. equal(A1::'a,B1) --> equal(ordinal_multiply(C1::'a,A1),ordinal_multiply(C1::'a,B1))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2318
\  (\\<forall>F1 G1 H1 I1. equal(F1::'a,G1) --> equal(recursion(F1::'a,H1,I1),recursion(G1::'a,H1,I1))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2319
\  (\\<forall>J1 L1 K1 M1. equal(J1::'a,K1) --> equal(recursion(L1::'a,J1,M1),recursion(L1::'a,K1,M1))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2320
\  (\\<forall>N1 P1 Q1 O1. equal(N1::'a,O1) --> equal(recursion(P1::'a,Q1,N1),recursion(P1::'a,Q1,O1))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2321
\  (\\<forall>R1 S1. equal(R1::'a,S1) --> equal(recursion_equation_functions(R1),recursion_equation_functions(S1))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2322
\  (\\<forall>T1 U1. equal(T1::'a,U1) --> equal(rest_of(T1),rest_of(U1))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2323
\  (\\<forall>V1 W1 X1 Y1. equal(V1::'a,W1) --> equal(segment(V1::'a,X1,Y1),segment(W1::'a,X1,Y1))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2324
\  (\\<forall>Z1 B2 A2 C2. equal(Z1::'a,A2) --> equal(segment(B2::'a,Z1,C2),segment(B2::'a,A2,C2))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2325
\  (\\<forall>D2 F2 G2 E2. equal(D2::'a,E2) --> equal(segment(F2::'a,G2,D2),segment(F2::'a,G2,E2))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2326
\  (\\<forall>H2 I2. equal(H2::'a,I2) --> equal(symmetrization_of(H2),symmetrization_of(I2))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2327
\  (\\<forall>J2 K2 L2. equal(J2::'a,K2) & asymmetric(J2::'a,L2) --> asymmetric(K2::'a,L2)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2328
\  (\\<forall>M2 O2 N2. equal(M2::'a,N2) & asymmetric(O2::'a,M2) --> asymmetric(O2::'a,N2)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2329
\  (\\<forall>P2 Q2 R2. equal(P2::'a,Q2) & connected(P2::'a,R2) --> connected(Q2::'a,R2)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2330
\  (\\<forall>S2 U2 T2. equal(S2::'a,T2) & connected(U2::'a,S2) --> connected(U2::'a,T2)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2331
\  (\\<forall>V2 W2 X2. equal(V2::'a,W2) & irreflexive(V2::'a,X2) --> irreflexive(W2::'a,X2)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2332
\  (\\<forall>Y2 A3 Z2. equal(Y2::'a,Z2) & irreflexive(A3::'a,Y2) --> irreflexive(A3::'a,Z2)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2333
\  (\\<forall>B3 C3 D3 E3. equal(B3::'a,C3) & section(B3::'a,D3,E3) --> section(C3::'a,D3,E3)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2334
\  (\\<forall>F3 H3 G3 I3. equal(F3::'a,G3) & section(H3::'a,F3,I3) --> section(H3::'a,G3,I3)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2335
\  (\\<forall>J3 L3 M3 K3. equal(J3::'a,K3) & section(L3::'a,M3,J3) --> section(L3::'a,M3,K3)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2336
\  (\\<forall>N3 O3 P3. equal(N3::'a,O3) & transitive(N3::'a,P3) --> transitive(O3::'a,P3)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2337
\  (\\<forall>Q3 S3 R3. equal(Q3::'a,R3) & transitive(S3::'a,Q3) --> transitive(S3::'a,R3)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2338
\  (\\<forall>T3 U3 V3. equal(T3::'a,U3) & well_ordering(T3::'a,V3) --> well_ordering(U3::'a,V3)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2339
\  (\\<forall>W3 Y3 X3. equal(W3::'a,X3) & well_ordering(Y3::'a,W3) --> well_ordering(Y3::'a,X3)) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2340
\  (~function(z)) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2341
\  (~equal(recursion_equation_functions(z),null_class)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2342
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2343
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2344
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2345
(*4868 inferences so far.  Searching to depth 12.  4.3 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2346
val PLA002_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2347
 ("(\\<forall>Situation1 Situation2. warm(Situation1) | cold(Situation2)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2348
\  (\\<forall>Situation. at(a::'a,Situation) --> at(b::'a,walk(b::'a,Situation))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2349
\  (\\<forall>Situation. at(a::'a,Situation) --> at(b::'a,drive(b::'a,Situation))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2350
\  (\\<forall>Situation. at(b::'a,Situation) --> at(a::'a,walk(a::'a,Situation))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2351
\  (\\<forall>Situation. at(b::'a,Situation) --> at(a::'a,drive(a::'a,Situation))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2352
\  (\\<forall>Situation. cold(Situation) & at(b::'a,Situation) --> at(c::'a,skate(c::'a,Situation))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2353
\  (\\<forall>Situation. cold(Situation) & at(c::'a,Situation) --> at(b::'a,skate(b::'a,Situation))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2354
\  (\\<forall>Situation. warm(Situation) & at(b::'a,Situation) --> at(d::'a,climb(d::'a,Situation))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2355
\  (\\<forall>Situation. warm(Situation) & at(d::'a,Situation) --> at(b::'a,climb(b::'a,Situation))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2356
\  (\\<forall>Situation. at(c::'a,Situation) --> at(d::'a,go(d::'a,Situation))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2357
\  (\\<forall>Situation. at(d::'a,Situation) --> at(c::'a,go(c::'a,Situation))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2358
\  (\\<forall>Situation. at(c::'a,Situation) --> at(e::'a,go(e::'a,Situation))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2359
\  (\\<forall>Situation. at(e::'a,Situation) --> at(c::'a,go(c::'a,Situation))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2360
\  (\\<forall>Situation. at(d::'a,Situation) --> at(f::'a,go(f::'a,Situation))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2361
\  (\\<forall>Situation. at(f::'a,Situation) --> at(d::'a,go(d::'a,Situation))) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2362
\  (at(f::'a,s0)) & \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2363
\  (\\<forall>S'. ~at(a::'a,S')) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2364
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2365
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2366
(*190 inferences so far.  Searching to depth 10.  0.6 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2367
val PLA006_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2368
 ("(\\<forall>X Y State. holds(X::'a,State) & holds(Y::'a,State) --> holds(and'(X::'a,Y),State)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2369
\  (\\<forall>State X. holds(EMPTY::'a,State) & holds(clear(X),State) & differ(X::'a,table) --> holds(holding(X),do(pickup(X),State))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2370
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2371
\  (\\<forall>Y State X Z. holds(on(X::'a,Y),State) & differ(X::'a,Z) --> holds(on(X::'a,Y),do(pickup(Z),State))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2372
\  (\\<forall>State X Z. holds(clear(X),State) & differ(X::'a,Z) --> holds(clear(X),do(pickup(Z),State))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2373
\  (\\<forall>X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(EMPTY::'a,do(putdown(X::'a,Y),State))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2374
\  (\\<forall>X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(on(X::'a,Y),do(putdown(X::'a,Y),State))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2375
\  (\\<forall>X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(clear(X),do(putdown(X::'a,Y),State))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2376
\  (\\<forall>Z W X Y State. holds(on(X::'a,Y),State) --> holds(on(X::'a,Y),do(putdown(Z::'a,W),State))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2377
\  (\\<forall>X State Z Y. holds(clear(Z),State) & differ(Z::'a,Y) --> holds(clear(Z),do(putdown(X::'a,Y),State))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2378
\  (\\<forall>Y X. differ(Y::'a,X) --> differ(X::'a,Y)) &       \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2379
\  (differ(a::'a,b)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2380
\  (differ(a::'a,c)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2381
\  (differ(a::'a,d)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2382
\  (differ(a::'a,table)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2383
\  (differ(b::'a,c)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2384
\  (differ(b::'a,d)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2385
\  (differ(b::'a,table)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2386
\  (differ(c::'a,d)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2387
\  (differ(c::'a,table)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2388
\  (differ(d::'a,table)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2389
\  (holds(on(a::'a,table),s0)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2390
\  (holds(on(b::'a,table),s0)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2391
\  (holds(on(c::'a,d),s0)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2392
\  (holds(on(d::'a,table),s0)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2393
\  (holds(clear(a),s0)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2394
\  (holds(clear(b),s0)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2395
\  (holds(clear(c),s0)) &       \
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2396
\  (holds(EMPTY::'a,s0)) &  \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2397
\  (\\<forall>State. holds(clear(table),State)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2398
\  (\\<forall>State. ~holds(on(c::'a,table),State)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2399
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2400
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2401
(*190 inferences so far.  Searching to depth 10.  0.5 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2402
val PLA017_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2403
 ("(\\<forall>X Y State. holds(X::'a,State) & holds(Y::'a,State) --> holds(and'(X::'a,Y),State)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2404
\  (\\<forall>State X. holds(EMPTY::'a,State) & holds(clear(X),State) & differ(X::'a,table) --> holds(holding(X),do(pickup(X),State))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2405
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2406
\  (\\<forall>Y State X Z. holds(on(X::'a,Y),State) & differ(X::'a,Z) --> holds(on(X::'a,Y),do(pickup(Z),State))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2407
\  (\\<forall>State X Z. holds(clear(X),State) & differ(X::'a,Z) --> holds(clear(X),do(pickup(Z),State))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2408
\  (\\<forall>X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(EMPTY::'a,do(putdown(X::'a,Y),State))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2409
\  (\\<forall>X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(on(X::'a,Y),do(putdown(X::'a,Y),State))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2410
\  (\\<forall>X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(clear(X),do(putdown(X::'a,Y),State))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2411
\  (\\<forall>Z W X Y State. holds(on(X::'a,Y),State) --> holds(on(X::'a,Y),do(putdown(Z::'a,W),State))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2412
\  (\\<forall>X State Z Y. holds(clear(Z),State) & differ(Z::'a,Y) --> holds(clear(Z),do(putdown(X::'a,Y),State))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2413
\  (\\<forall>Y X. differ(Y::'a,X) --> differ(X::'a,Y)) &       \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2414
\  (differ(a::'a,b)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2415
\  (differ(a::'a,c)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2416
\  (differ(a::'a,d)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2417
\  (differ(a::'a,table)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2418
\  (differ(b::'a,c)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2419
\  (differ(b::'a,d)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2420
\  (differ(b::'a,table)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2421
\  (differ(c::'a,d)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2422
\  (differ(c::'a,table)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2423
\  (differ(d::'a,table)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2424
\  (holds(on(a::'a,table),s0)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2425
\  (holds(on(b::'a,table),s0)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2426
\  (holds(on(c::'a,d),s0)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2427
\  (holds(on(d::'a,table),s0)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2428
\  (holds(clear(a),s0)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2429
\  (holds(clear(b),s0)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2430
\  (holds(clear(c),s0)) &       \
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2431
\  (holds(EMPTY::'a,s0)) &  \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2432
\  (\\<forall>State. holds(clear(table),State)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2433
\  (\\<forall>State. ~holds(on(a::'a,c),State)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2434
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2435
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2436
(*13732 inferences so far.  Searching to depth 16.  9.8 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2437
val PLA022_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2438
 ("(\\<forall>X Y State. holds(X::'a,State) & holds(Y::'a,State) --> holds(and'(X::'a,Y),State)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2439
\  (\\<forall>State X. holds(EMPTY::'a,State) & holds(clear(X),State) & differ(X::'a,table) --> holds(holding(X),do(pickup(X),State))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2440
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2441
\  (\\<forall>Y State X Z. holds(on(X::'a,Y),State) & differ(X::'a,Z) --> holds(on(X::'a,Y),do(pickup(Z),State))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2442
\  (\\<forall>State X Z. holds(clear(X),State) & differ(X::'a,Z) --> holds(clear(X),do(pickup(Z),State))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2443
\  (\\<forall>X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(EMPTY::'a,do(putdown(X::'a,Y),State))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2444
\  (\\<forall>X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(on(X::'a,Y),do(putdown(X::'a,Y),State))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2445
\  (\\<forall>X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(clear(X),do(putdown(X::'a,Y),State))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2446
\  (\\<forall>Z W X Y State. holds(on(X::'a,Y),State) --> holds(on(X::'a,Y),do(putdown(Z::'a,W),State))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2447
\  (\\<forall>X State Z Y. holds(clear(Z),State) & differ(Z::'a,Y) --> holds(clear(Z),do(putdown(X::'a,Y),State))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2448
\  (\\<forall>Y X. differ(Y::'a,X) --> differ(X::'a,Y)) &       \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2449
\  (differ(a::'a,b)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2450
\  (differ(a::'a,c)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2451
\  (differ(a::'a,d)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2452
\  (differ(a::'a,table)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2453
\  (differ(b::'a,c)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2454
\  (differ(b::'a,d)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2455
\  (differ(b::'a,table)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2456
\  (differ(c::'a,d)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2457
\  (differ(c::'a,table)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2458
\  (differ(d::'a,table)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2459
\  (holds(on(a::'a,table),s0)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2460
\  (holds(on(b::'a,table),s0)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2461
\  (holds(on(c::'a,d),s0)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2462
\  (holds(on(d::'a,table),s0)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2463
\  (holds(clear(a),s0)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2464
\  (holds(clear(b),s0)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2465
\  (holds(clear(c),s0)) &       \
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2466
\  (holds(EMPTY::'a,s0)) &  \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2467
\  (\\<forall>State. holds(clear(table),State)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2468
\  (\\<forall>State. ~holds(and'(on(c::'a,d),on(a::'a,c)),State)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2469
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2470
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2471
(*217 inferences so far.  Searching to depth 13.  0.7 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2472
val PLA022_2 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2473
 ("(\\<forall>X Y State. holds(X::'a,State) & holds(Y::'a,State) --> holds(and'(X::'a,Y),State)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2474
\  (\\<forall>State X. holds(EMPTY::'a,State) & holds(clear(X),State) & differ(X::'a,table) --> holds(holding(X),do(pickup(X),State))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2475
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2476
\  (\\<forall>Y State X Z. holds(on(X::'a,Y),State) & differ(X::'a,Z) --> holds(on(X::'a,Y),do(pickup(Z),State))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2477
\  (\\<forall>State X Z. holds(clear(X),State) & differ(X::'a,Z) --> holds(clear(X),do(pickup(Z),State))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2478
\  (\\<forall>X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(EMPTY::'a,do(putdown(X::'a,Y),State))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2479
\  (\\<forall>X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(on(X::'a,Y),do(putdown(X::'a,Y),State))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2480
\  (\\<forall>X Y State. holds(holding(X),State) & holds(clear(Y),State) --> holds(clear(X),do(putdown(X::'a,Y),State))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2481
\  (\\<forall>Z W X Y State. holds(on(X::'a,Y),State) --> holds(on(X::'a,Y),do(putdown(Z::'a,W),State))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2482
\  (\\<forall>X State Z Y. holds(clear(Z),State) & differ(Z::'a,Y) --> holds(clear(Z),do(putdown(X::'a,Y),State))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2483
\  (\\<forall>Y X. differ(Y::'a,X) --> differ(X::'a,Y)) &       \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2484
\  (differ(a::'a,b)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2485
\  (differ(a::'a,c)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2486
\  (differ(a::'a,d)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2487
\  (differ(a::'a,table)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2488
\  (differ(b::'a,c)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2489
\  (differ(b::'a,d)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2490
\  (differ(b::'a,table)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2491
\  (differ(c::'a,d)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2492
\  (differ(c::'a,table)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2493
\  (differ(d::'a,table)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2494
\  (holds(on(a::'a,table),s0)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2495
\  (holds(on(b::'a,table),s0)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2496
\  (holds(on(c::'a,d),s0)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2497
\  (holds(on(d::'a,table),s0)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2498
\  (holds(clear(a),s0)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2499
\  (holds(clear(b),s0)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2500
\  (holds(clear(c),s0)) &       \
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2501
\  (holds(EMPTY::'a,s0)) &  \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2502
\  (\\<forall>State. holds(clear(table),State)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2503
\  (\\<forall>State. ~holds(and'(on(a::'a,c),on(c::'a,d)),State)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2504
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2505
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2506
(*948 inferences so far.  Searching to depth 18.  1.1 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2507
val PRV001_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2508
 ("(\\<forall>X Y Z. q1(X::'a,Y,Z) & less_or_equal(X::'a,Y) --> q2(X::'a,Y,Z)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2509
\  (\\<forall>X Y Z. q1(X::'a,Y,Z) --> less_or_equal(X::'a,Y) | q3(X::'a,Y,Z)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2510
\  (\\<forall>Z X Y. q2(X::'a,Y,Z) --> q4(X::'a,Y,Y)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2511
\  (\\<forall>Z Y X. q3(X::'a,Y,Z) --> q4(X::'a,Y,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2512
\  (\\<forall>X. less_or_equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2513
\  (\\<forall>X Y. less_or_equal(X::'a,Y) & less_or_equal(Y::'a,X) --> equal(X::'a,Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2514
\  (\\<forall>Y X Z. less_or_equal(X::'a,Y) & less_or_equal(Y::'a,Z) --> less_or_equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2515
\  (\\<forall>Y X. less_or_equal(X::'a,Y) | less_or_equal(Y::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2516
\  (\\<forall>X Y. equal(X::'a,Y) --> less_or_equal(X::'a,Y)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2517
\  (\\<forall>X Y Z. equal(X::'a,Y) & less_or_equal(X::'a,Z) --> less_or_equal(Y::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2518
\  (\\<forall>X Z Y. equal(X::'a,Y) & less_or_equal(Z::'a,X) --> less_or_equal(Z::'a,Y)) &  \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2519
\  (q1(a::'a,b,c)) &        \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2520
\  (\\<forall>W. ~(q4(a::'a,b,W) & less_or_equal(a::'a,W) & less_or_equal(b::'a,W) & less_or_equal(W::'a,a))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2521
\  (\\<forall>W. ~(q4(a::'a,b,W) & less_or_equal(a::'a,W) & less_or_equal(b::'a,W) & less_or_equal(W::'a,b))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2522
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2523
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2524
(*21 inferences so far.  Searching to depth 5.  0.4 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2525
val PRV003_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2526
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2527
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2528
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2529
\  (\\<forall>X. equal(predecessor(successor(X)),X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2530
\  (\\<forall>X. equal(successor(predecessor(X)),X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2531
\  (\\<forall>X Y. equal(predecessor(X),predecessor(Y)) --> equal(X::'a,Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2532
\  (\\<forall>X Y. equal(successor(X),successor(Y)) --> equal(X::'a,Y)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2533
\  (\\<forall>X. LESS_THAN(predecessor(X),X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2534
\  (\\<forall>X. LESS_THAN(X::'a,successor(X))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2535
\  (\\<forall>X Y Z. LESS_THAN(X::'a,Y) & LESS_THAN(Y::'a,Z) --> LESS_THAN(X::'a,Z)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2536
\  (\\<forall>X Y. LESS_THAN(X::'a,Y) | LESS_THAN(Y::'a,X) | equal(X::'a,Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2537
\  (\\<forall>X. ~LESS_THAN(X::'a,X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2538
\  (\\<forall>Y X. ~(LESS_THAN(X::'a,Y) & LESS_THAN(Y::'a,X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2539
\  (\\<forall>Y X Z. equal(X::'a,Y) & LESS_THAN(X::'a,Z) --> LESS_THAN(Y::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2540
\  (\\<forall>Y Z X. equal(X::'a,Y) & LESS_THAN(Z::'a,X) --> LESS_THAN(Z::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2541
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(predecessor(X),predecessor(Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2542
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(successor(X),successor(Y))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2543
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(a(X),a(Y))) &   \
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2544
\  (~LESS_THAN(n::'a,j)) &  \
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2545
\  (LESS_THAN(k::'a,j)) &   \
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2546
\  (~LESS_THAN(k::'a,i)) &  \
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2547
\  (LESS_THAN(i::'a,n)) &   \
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2548
\  (LESS_THAN(a(j),a(k))) &     \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2549
\  (\\<forall>X. LESS_THAN(X::'a,j) & LESS_THAN(a(X),a(k)) --> LESS_THAN(X::'a,i)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2550
\  (\\<forall>X. LESS_THAN(one::'a,i) & LESS_THAN(a(X),a(predecessor(i))) --> LESS_THAN(X::'a,i) | LESS_THAN(n::'a,X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2551
\  (\\<forall>X. ~(LESS_THAN(one::'a,X) & LESS_THAN(X::'a,i) & LESS_THAN(a(X),a(predecessor(X))))) &    \
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2552
\  (LESS_THAN(j::'a,i)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2553
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2554
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2555
(*584 inferences so far.  Searching to depth 7.  1.1 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2556
val PRV005_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2557
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2558
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2559
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2560
\  (\\<forall>X. equal(predecessor(successor(X)),X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2561
\  (\\<forall>X. equal(successor(predecessor(X)),X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2562
\  (\\<forall>X Y. equal(predecessor(X),predecessor(Y)) --> equal(X::'a,Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2563
\  (\\<forall>X Y. equal(successor(X),successor(Y)) --> equal(X::'a,Y)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2564
\  (\\<forall>X. LESS_THAN(predecessor(X),X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2565
\  (\\<forall>X. LESS_THAN(X::'a,successor(X))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2566
\  (\\<forall>X Y Z. LESS_THAN(X::'a,Y) & LESS_THAN(Y::'a,Z) --> LESS_THAN(X::'a,Z)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2567
\  (\\<forall>X Y. LESS_THAN(X::'a,Y) | LESS_THAN(Y::'a,X) | equal(X::'a,Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2568
\  (\\<forall>X. ~LESS_THAN(X::'a,X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2569
\  (\\<forall>Y X. ~(LESS_THAN(X::'a,Y) & LESS_THAN(Y::'a,X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2570
\  (\\<forall>Y X Z. equal(X::'a,Y) & LESS_THAN(X::'a,Z) --> LESS_THAN(Y::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2571
\  (\\<forall>Y Z X. equal(X::'a,Y) & LESS_THAN(Z::'a,X) --> LESS_THAN(Z::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2572
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(predecessor(X),predecessor(Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2573
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(successor(X),successor(Y))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2574
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(a(X),a(Y))) &   \
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2575
\  (~LESS_THAN(n::'a,k)) &  \
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2576
\  (~LESS_THAN(k::'a,l)) &  \
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2577
\  (~LESS_THAN(k::'a,i)) &  \
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2578
\  (LESS_THAN(l::'a,n)) &   \
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2579
\  (LESS_THAN(one::'a,l)) & \
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2580
\  (LESS_THAN(a(k),a(predecessor(l)))) &        \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2581
\  (\\<forall>X. LESS_THAN(X::'a,successor(n)) & LESS_THAN(a(X),a(k)) --> LESS_THAN(X::'a,l)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2582
\  (\\<forall>X. LESS_THAN(one::'a,l) & LESS_THAN(a(X),a(predecessor(l))) --> LESS_THAN(X::'a,l) | LESS_THAN(n::'a,X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2583
\  (\\<forall>X. ~(LESS_THAN(one::'a,X) & LESS_THAN(X::'a,l) & LESS_THAN(a(X),a(predecessor(X))))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2584
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2585
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2586
(*2343 inferences so far.  Searching to depth 8.  3.5 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2587
val PRV006_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2588
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2589
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2590
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2591
\  (\\<forall>X. equal(predecessor(successor(X)),X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2592
\  (\\<forall>X. equal(successor(predecessor(X)),X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2593
\  (\\<forall>X Y. equal(predecessor(X),predecessor(Y)) --> equal(X::'a,Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2594
\  (\\<forall>X Y. equal(successor(X),successor(Y)) --> equal(X::'a,Y)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2595
\  (\\<forall>X. LESS_THAN(predecessor(X),X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2596
\  (\\<forall>X. LESS_THAN(X::'a,successor(X))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2597
\  (\\<forall>X Y Z. LESS_THAN(X::'a,Y) & LESS_THAN(Y::'a,Z) --> LESS_THAN(X::'a,Z)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2598
\  (\\<forall>X Y. LESS_THAN(X::'a,Y) | LESS_THAN(Y::'a,X) | equal(X::'a,Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2599
\  (\\<forall>X. ~LESS_THAN(X::'a,X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2600
\  (\\<forall>Y X. ~(LESS_THAN(X::'a,Y) & LESS_THAN(Y::'a,X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2601
\  (\\<forall>Y X Z. equal(X::'a,Y) & LESS_THAN(X::'a,Z) --> LESS_THAN(Y::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2602
\  (\\<forall>Y Z X. equal(X::'a,Y) & LESS_THAN(Z::'a,X) --> LESS_THAN(Z::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2603
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(predecessor(X),predecessor(Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2604
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(successor(X),successor(Y))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2605
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(a(X),a(Y))) &   \
11451
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2606
\  (~LESS_THAN(n::'a,m)) &  \
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2607
\  (LESS_THAN(i::'a,m)) &   \
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2608
\  (LESS_THAN(i::'a,n)) &   \
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2609
\  (~LESS_THAN(i::'a,one)) &        \
8abfb4f7bd02 partial restructuring to reduce dependence on Axiom of Choice
paulson
parents: 10440
diff changeset
  2610
\  (LESS_THAN(a(i),a(m))) &     \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2611
\  (\\<forall>X. LESS_THAN(X::'a,successor(n)) & LESS_THAN(a(X),a(m)) --> LESS_THAN(X::'a,i)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2612
\  (\\<forall>X. LESS_THAN(one::'a,i) & LESS_THAN(a(X),a(predecessor(i))) --> LESS_THAN(X::'a,i) | LESS_THAN(n::'a,X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2613
\  (\\<forall>X. ~(LESS_THAN(one::'a,X) & LESS_THAN(X::'a,i) & LESS_THAN(a(X),a(predecessor(X))))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2614
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2615
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2616
(*86 inferences so far.  Searching to depth 14.  0.1 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2617
val PRV009_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2618
 ("(\\<forall>Y X. less_or_equal(X::'a,Y) | less(Y::'a,X)) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2619
\  (less(j::'a,i)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2620
\  (less_or_equal(m::'a,p)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2621
\  (less_or_equal(p::'a,q)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2622
\  (less_or_equal(q::'a,n)) &       \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2623
\  (\\<forall>X Y. less_or_equal(m::'a,X) & less(X::'a,i) & less(j::'a,Y) & less_or_equal(Y::'a,n) --> less_or_equal(a(X),a(Y))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2624
\  (\\<forall>X Y. less_or_equal(m::'a,X) & less_or_equal(X::'a,Y) & less_or_equal(Y::'a,j) --> less_or_equal(a(X),a(Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2625
\  (\\<forall>X Y. less_or_equal(i::'a,X) & less_or_equal(X::'a,Y) & less_or_equal(Y::'a,n) --> less_or_equal(a(X),a(Y))) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2626
\  (~less_or_equal(a(p),a(q))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2627
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2628
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2629
(*222 inferences so far.  Searching to depth 8.  0.4 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2630
val PUZ012_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2631
 ("(\\<forall>X. equal_fruits(X::'a,X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2632
\  (\\<forall>X. equal_boxes(X::'a,X)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2633
\  (\\<forall>X Y. ~(label(X::'a,Y) & contains(X::'a,Y))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2634
\  (\\<forall>X. contains(boxa::'a,X) | contains(boxb::'a,X) | contains(boxc::'a,X)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2635
\  (\\<forall>X. contains(X::'a,apples) | contains(X::'a,bananas) | contains(X::'a,oranges)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2636
\  (\\<forall>X Y Z. contains(X::'a,Y) & contains(X::'a,Z) --> equal_fruits(Y::'a,Z)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2637
\  (\\<forall>Y X Z. contains(X::'a,Y) & contains(Z::'a,Y) --> equal_boxes(X::'a,Z)) &      \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2638
\  (~equal_boxes(boxa::'a,boxb)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2639
\  (~equal_boxes(boxb::'a,boxc)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2640
\  (~equal_boxes(boxa::'a,boxc)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2641
\  (~equal_fruits(apples::'a,bananas)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2642
\  (~equal_fruits(bananas::'a,oranges)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2643
\  (~equal_fruits(apples::'a,oranges)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2644
\  (label(boxa::'a,apples)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2645
\  (label(boxb::'a,oranges)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2646
\  (label(boxc::'a,bananas)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2647
\  (contains(boxb::'a,apples)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2648
\  (~(contains(boxa::'a,bananas) & contains(boxc::'a,oranges))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2649
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2650
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2651
(*35 inferences so far.  Searching to depth 5.  3.2 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2652
val PUZ020_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2653
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2654
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2655
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2656
\  (\\<forall>A B. equal(A::'a,B) --> equal(statement_by(A),statement_by(B))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2657
\  (\\<forall>X. person(X) --> knight(X) | knave(X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2658
\  (\\<forall>X. ~(person(X) & knight(X) & knave(X))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2659
\  (\\<forall>X Y. says(X::'a,Y) & a_truth(Y) --> a_truth(Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2660
\  (\\<forall>X Y. ~(says(X::'a,Y) & equal(X::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2661
\  (\\<forall>Y X. says(X::'a,Y) --> equal(Y::'a,statement_by(X))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2662
\  (\\<forall>X Y. ~(person(X) & equal(X::'a,statement_by(Y)))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2663
\  (\\<forall>X. person(X) & a_truth(statement_by(X)) --> knight(X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2664
\  (\\<forall>X. person(X) --> a_truth(statement_by(X)) | knave(X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2665
\  (\\<forall>X Y. equal(X::'a,Y) & knight(X) --> knight(Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2666
\  (\\<forall>X Y. equal(X::'a,Y) & knave(X) --> knave(Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2667
\  (\\<forall>X Y. equal(X::'a,Y) & person(X) --> person(Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2668
\  (\\<forall>X Y Z. equal(X::'a,Y) & says(X::'a,Z) --> says(Y::'a,Z)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2669
\  (\\<forall>X Z Y. equal(X::'a,Y) & says(Z::'a,X) --> says(Z::'a,Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2670
\  (\\<forall>X Y. equal(X::'a,Y) & a_truth(X) --> a_truth(Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2671
\  (\\<forall>X Y. knight(X) & says(X::'a,Y) --> a_truth(Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2672
\  (\\<forall>X Y. ~(knave(X) & says(X::'a,Y) & a_truth(Y))) &      \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2673
\  (person(husband)) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2674
\  (person(wife)) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2675
\  (~equal(husband::'a,wife)) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2676
\  (says(husband::'a,statement_by(husband))) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2677
\  (a_truth(statement_by(husband)) & knight(husband) --> knight(wife)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2678
\  (knight(husband) --> a_truth(statement_by(husband))) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2679
\  (a_truth(statement_by(husband)) | knight(wife)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2680
\  (knight(wife) --> a_truth(statement_by(husband))) &  \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2681
\  (~knight(husband)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2682
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2683
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2684
(*121806 inferences so far.  Searching to depth 17.  63.0 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2685
val PUZ025_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2686
 ("(\\<forall>X. a_truth(truthteller(X)) | a_truth(liar(X))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2687
\  (\\<forall>X. ~(a_truth(truthteller(X)) & a_truth(liar(X)))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2688
\  (\\<forall>Truthteller Statement. a_truth(truthteller(Truthteller)) & a_truth(says(Truthteller::'a,Statement)) --> a_truth(Statement)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2689
\  (\\<forall>Liar Statement. ~(a_truth(liar(Liar)) & a_truth(says(Liar::'a,Statement)) & a_truth(Statement))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2690
\  (\\<forall>Statement Truthteller. a_truth(Statement) & a_truth(says(Truthteller::'a,Statement)) --> a_truth(truthteller(Truthteller))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2691
\  (\\<forall>Statement Liar. a_truth(says(Liar::'a,Statement)) --> a_truth(Statement) | a_truth(liar(Liar))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2692
\  (\\<forall>Z X Y. people(X::'a,Y,Z) & a_truth(liar(X)) & a_truth(liar(Y)) --> a_truth(equal_type(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2693
\  (\\<forall>Z X Y. people(X::'a,Y,Z) & a_truth(truthteller(X)) & a_truth(truthteller(Y)) --> a_truth(equal_type(X::'a,Y))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2694
\  (\\<forall>X Y. a_truth(equal_type(X::'a,Y)) & a_truth(truthteller(X)) --> a_truth(truthteller(Y))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2695
\  (\\<forall>X Y. a_truth(equal_type(X::'a,Y)) & a_truth(liar(X)) --> a_truth(liar(Y))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2696
\  (\\<forall>X Y. a_truth(truthteller(X)) --> a_truth(equal_type(X::'a,Y)) | a_truth(liar(Y))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2697
\  (\\<forall>X Y. a_truth(liar(X)) --> a_truth(equal_type(X::'a,Y)) | a_truth(truthteller(Y))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2698
\  (\\<forall>Y X. a_truth(equal_type(X::'a,Y)) --> a_truth(equal_type(Y::'a,X))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2699
\  (\\<forall>X Y. ask_1_if_2(X::'a,Y) & a_truth(truthteller(X)) & a_truth(Y) --> answer(yes)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2700
\  (\\<forall>X Y. ask_1_if_2(X::'a,Y) & a_truth(truthteller(X)) --> a_truth(Y) | answer(no)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2701
\  (\\<forall>X Y. ask_1_if_2(X::'a,Y) & a_truth(liar(X)) & a_truth(Y) --> answer(no)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2702
\  (\\<forall>X Y. ask_1_if_2(X::'a,Y) & a_truth(liar(X)) --> a_truth(Y) | answer(yes)) &  \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2703
\  (people(b::'a,c,a)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2704
\  (people(a::'a,b,a)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2705
\  (people(a::'a,c,b)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2706
\  (people(c::'a,b,a)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2707
\  (a_truth(says(a::'a,equal_type(b::'a,c)))) & \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2708
\  (ask_1_if_2(c::'a,equal_type(a::'a,b))) &    \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2709
\  (\\<forall>Answer. ~answer(Answer)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2710
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2711
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2712
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2713
(*621 inferences so far.  Searching to depth 18.  0.2 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2714
val PUZ029_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2715
 ("(\\<forall>X. dances_on_tightropes(X) | eats_pennybuns(X) | old(X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2716
\  (\\<forall>X. pig(X) & liable_to_giddiness(X) --> treated_with_respect(X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2717
\  (\\<forall>X. wise(X) & balloonist(X) --> has_umbrella(X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2718
\  (\\<forall>X. ~(looks_ridiculous(X) & eats_pennybuns(X) & eats_lunch_in_public(X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2719
\  (\\<forall>X. balloonist(X) & young(X) --> liable_to_giddiness(X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2720
\  (\\<forall>X. fat(X) & looks_ridiculous(X) --> dances_on_tightropes(X) | eats_lunch_in_public(X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2721
\  (\\<forall>X. ~(liable_to_giddiness(X) & wise(X) & dances_on_tightropes(X))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2722
\  (\\<forall>X. pig(X) & has_umbrella(X) --> looks_ridiculous(X)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2723
\  (\\<forall>X. treated_with_respect(X) --> dances_on_tightropes(X) | fat(X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2724
\  (\\<forall>X. young(X) | old(X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2725
\  (\\<forall>X. ~(young(X) & old(X))) &        \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2726
\  (wise(piggy)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2727
\  (young(piggy)) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2728
\  (pig(piggy)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2729
\  (balloonist(piggy)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2730
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2731
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2732
(*93620 inferences so far.  Searching to depth 24.  65.9 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2733
val RNG001_3 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2734
 ("(\\<forall>X. sum(additive_identity::'a,X,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2735
\  (\\<forall>X. sum(additive_inverse(X),X,additive_identity)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2736
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2737
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2738
\  (\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2739
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2740
\  (\\<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)) &        \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2741
\  (~product(a::'a,additive_identity,additive_identity)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2742
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2743
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2744
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2745
(****************SLOW
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2746
3057170 inferences so far.  Searching to depth 16.  No proof after 45 minutes.
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2747
val RNG001_5 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2748
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2749
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2750
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2751
\  (\\<forall>X. sum(additive_identity::'a,X,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2752
\  (\\<forall>X. sum(X::'a,additive_identity,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2753
\  (\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2754
\  (\\<forall>X Y. sum(X::'a,Y,add(X::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2755
\  (\\<forall>X. sum(additive_inverse(X),X,additive_identity)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2756
\  (\\<forall>X. sum(X::'a,additive_inverse(X),additive_identity)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2757
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2758
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2759
\  (\\<forall>Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2760
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2761
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2762
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2763
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2764
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2765
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2766
\  (\\<forall>X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2767
\  (\\<forall>X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2768
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(additive_inverse(X),additive_inverse(Y))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2769
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(add(X::'a,W),add(Y::'a,W))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2770
\  (\\<forall>X Y W Z. equal(X::'a,Y) & sum(X::'a,W,Z) --> sum(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2771
\  (\\<forall>X W Y Z. equal(X::'a,Y) & sum(W::'a,X,Z) --> sum(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2772
\  (\\<forall>X W Z Y. equal(X::'a,Y) & sum(W::'a,Z,X) --> sum(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2773
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2774
\  (\\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2775
\  (\\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2776
\  (\\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) &        \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2777
\  (~product(a::'a,additive_identity,additive_identity)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2778
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2779
****************)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2780
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2781
(*0 inferences so far.  Searching to depth 0.  0.5 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2782
val RNG011_5 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2783
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2784
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2785
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2786
\  (\\<forall>A B C. equal(A::'a,B) --> equal(add(A::'a,C),add(B::'a,C))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2787
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(add(F'::'a,D),add(F'::'a,E))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2788
\  (\\<forall>G H. equal(G::'a,H) --> equal(additive_inverse(G),additive_inverse(H))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2789
\  (\\<forall>I' J K'. equal(I'::'a,J) --> equal(multiply(I'::'a,K'),multiply(J::'a,K'))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2790
\  (\\<forall>L N M. equal(L::'a,M) --> equal(multiply(N::'a,L),multiply(N::'a,M))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2791
\  (\\<forall>A B C D. equal(A::'a,B) --> equal(associator(A::'a,C,D),associator(B::'a,C,D))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2792
\  (\\<forall>E G F' H. equal(E::'a,F') --> equal(associator(G::'a,E,H),associator(G::'a,F',H))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2793
\  (\\<forall>I' K' L J. equal(I'::'a,J) --> equal(associator(K'::'a,L,I'),associator(K'::'a,L,J))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2794
\  (\\<forall>M N O_. equal(M::'a,N) --> equal(commutator(M::'a,O_),commutator(N::'a,O_))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2795
\  (\\<forall>P R Q. equal(P::'a,Q) --> equal(commutator(R::'a,P),commutator(R::'a,Q))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2796
\  (\\<forall>Y X. equal(add(X::'a,Y),add(Y::'a,X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2797
\  (\\<forall>X Y Z. equal(add(add(X::'a,Y),Z),add(X::'a,add(Y::'a,Z)))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2798
\  (\\<forall>X. equal(add(X::'a,additive_identity),X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2799
\  (\\<forall>X. equal(add(additive_identity::'a,X),X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2800
\  (\\<forall>X. equal(add(X::'a,additive_inverse(X)),additive_identity)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2801
\  (\\<forall>X. equal(add(additive_inverse(X),X),additive_identity)) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2802
\  (equal(additive_inverse(additive_identity),additive_identity)) &     \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2803
\  (\\<forall>X Y. equal(add(X::'a,add(additive_inverse(X),Y)),Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2804
\  (\\<forall>X Y. equal(additive_inverse(add(X::'a,Y)),add(additive_inverse(X),additive_inverse(Y)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2805
\  (\\<forall>X. equal(additive_inverse(additive_inverse(X)),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2806
\  (\\<forall>X. equal(multiply(X::'a,additive_identity),additive_identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2807
\  (\\<forall>X. equal(multiply(additive_identity::'a,X),additive_identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2808
\  (\\<forall>X Y. equal(multiply(additive_inverse(X),additive_inverse(Y)),multiply(X::'a,Y))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2809
\  (\\<forall>X Y. equal(multiply(X::'a,additive_inverse(Y)),additive_inverse(multiply(X::'a,Y)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2810
\  (\\<forall>X Y. equal(multiply(additive_inverse(X),Y),additive_inverse(multiply(X::'a,Y)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2811
\  (\\<forall>Y X Z. equal(multiply(X::'a,add(Y::'a,Z)),add(multiply(X::'a,Y),multiply(X::'a,Z)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2812
\  (\\<forall>X Y Z. equal(multiply(add(X::'a,Y),Z),add(multiply(X::'a,Z),multiply(Y::'a,Z)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2813
\  (\\<forall>X Y. equal(multiply(multiply(X::'a,Y),Y),multiply(X::'a,multiply(Y::'a,Y)))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2814
\  (\\<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)))))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2815
\  (\\<forall>X Y. equal(commutator(X::'a,Y),add(multiply(Y::'a,X),additive_inverse(multiply(X::'a,Y))))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2816
\  (\\<forall>X Y. equal(multiply(multiply(associator(X::'a,X,Y),X),associator(X::'a,X,Y)),additive_identity)) &        \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2817
\  (~equal(multiply(multiply(associator(a::'a,a,b),a),associator(a::'a,a,b)),additive_identity)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2818
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2819
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2820
(*202 inferences so far.  Searching to depth 8.  0.6 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2821
val RNG023_6 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2822
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2823
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2824
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2825
\  (\\<forall>Y X. equal(add(X::'a,Y),add(Y::'a,X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2826
\  (\\<forall>X Y Z. equal(add(X::'a,add(Y::'a,Z)),add(add(X::'a,Y),Z))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2827
\  (\\<forall>X. equal(add(additive_identity::'a,X),X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2828
\  (\\<forall>X. equal(add(X::'a,additive_identity),X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2829
\  (\\<forall>X. equal(multiply(additive_identity::'a,X),additive_identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2830
\  (\\<forall>X. equal(multiply(X::'a,additive_identity),additive_identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2831
\  (\\<forall>X. equal(add(additive_inverse(X),X),additive_identity)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2832
\  (\\<forall>X. equal(add(X::'a,additive_inverse(X)),additive_identity)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2833
\  (\\<forall>Y X Z. equal(multiply(X::'a,add(Y::'a,Z)),add(multiply(X::'a,Y),multiply(X::'a,Z)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2834
\  (\\<forall>X Y Z. equal(multiply(add(X::'a,Y),Z),add(multiply(X::'a,Z),multiply(Y::'a,Z)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2835
\  (\\<forall>X. equal(additive_inverse(additive_inverse(X)),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2836
\  (\\<forall>X Y. equal(multiply(multiply(X::'a,Y),Y),multiply(X::'a,multiply(Y::'a,Y)))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2837
\  (\\<forall>X Y. equal(multiply(multiply(X::'a,X),Y),multiply(X::'a,multiply(X::'a,Y)))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2838
\  (\\<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)))))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2839
\  (\\<forall>X Y. equal(commutator(X::'a,Y),add(multiply(Y::'a,X),additive_inverse(multiply(X::'a,Y))))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2840
\  (\\<forall>D E F'. equal(D::'a,E) --> equal(add(D::'a,F'),add(E::'a,F'))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2841
\  (\\<forall>G I' H. equal(G::'a,H) --> equal(add(I'::'a,G),add(I'::'a,H))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2842
\  (\\<forall>J K'. equal(J::'a,K') --> equal(additive_inverse(J),additive_inverse(K'))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2843
\  (\\<forall>L M N O_. equal(L::'a,M) --> equal(associator(L::'a,N,O_),associator(M::'a,N,O_))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2844
\  (\\<forall>P R Q S'. equal(P::'a,Q) --> equal(associator(R::'a,P,S'),associator(R::'a,Q,S'))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2845
\  (\\<forall>T' V W U. equal(T'::'a,U) --> equal(associator(V::'a,W,T'),associator(V::'a,W,U))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2846
\  (\\<forall>X Y Z. equal(X::'a,Y) --> equal(commutator(X::'a,Z),commutator(Y::'a,Z))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2847
\  (\\<forall>A1 C1 B1. equal(A1::'a,B1) --> equal(commutator(C1::'a,A1),commutator(C1::'a,B1))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2848
\  (\\<forall>D1 E1 F1. equal(D1::'a,E1) --> equal(multiply(D1::'a,F1),multiply(E1::'a,F1))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2849
\  (\\<forall>G1 I1 H1. equal(G1::'a,H1) --> equal(multiply(I1::'a,G1),multiply(I1::'a,H1))) &      \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2850
\  (~equal(associator(x::'a,x,y),additive_identity)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2851
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2852
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2853
(*0 inferences so far.  Searching to depth 0.  0.6 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2854
val RNG028_2 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2855
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2856
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2857
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2858
\  (\\<forall>X. equal(add(additive_identity::'a,X),X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2859
\  (\\<forall>X. equal(multiply(additive_identity::'a,X),additive_identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2860
\  (\\<forall>X. equal(multiply(X::'a,additive_identity),additive_identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2861
\  (\\<forall>X. equal(add(additive_inverse(X),X),additive_identity)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2862
\  (\\<forall>X Y. equal(additive_inverse(add(X::'a,Y)),add(additive_inverse(X),additive_inverse(Y)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2863
\  (\\<forall>X. equal(additive_inverse(additive_inverse(X)),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2864
\  (\\<forall>Y X Z. equal(multiply(X::'a,add(Y::'a,Z)),add(multiply(X::'a,Y),multiply(X::'a,Z)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2865
\  (\\<forall>X Y Z. equal(multiply(add(X::'a,Y),Z),add(multiply(X::'a,Z),multiply(Y::'a,Z)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2866
\  (\\<forall>X Y. equal(multiply(multiply(X::'a,Y),Y),multiply(X::'a,multiply(Y::'a,Y)))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2867
\  (\\<forall>X Y. equal(multiply(multiply(X::'a,X),Y),multiply(X::'a,multiply(X::'a,Y)))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2868
\  (\\<forall>X Y. equal(multiply(additive_inverse(X),Y),additive_inverse(multiply(X::'a,Y)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2869
\  (\\<forall>X Y. equal(multiply(X::'a,additive_inverse(Y)),additive_inverse(multiply(X::'a,Y)))) &    \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2870
\  (equal(additive_inverse(additive_identity),additive_identity)) &     \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2871
\  (\\<forall>Y X. equal(add(X::'a,Y),add(Y::'a,X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2872
\  (\\<forall>X Y Z. equal(add(X::'a,add(Y::'a,Z)),add(add(X::'a,Y),Z))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2873
\  (\\<forall>Z X Y. equal(add(X::'a,Z),add(Y::'a,Z)) --> equal(X::'a,Y)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2874
\  (\\<forall>Z X Y. equal(add(Z::'a,X),add(Z::'a,Y)) --> equal(X::'a,Y)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2875
\  (\\<forall>D E F'. equal(D::'a,E) --> equal(add(D::'a,F'),add(E::'a,F'))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2876
\  (\\<forall>G I' H. equal(G::'a,H) --> equal(add(I'::'a,G),add(I'::'a,H))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2877
\  (\\<forall>J K'. equal(J::'a,K') --> equal(additive_inverse(J),additive_inverse(K'))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2878
\  (\\<forall>D1 E1 F1. equal(D1::'a,E1) --> equal(multiply(D1::'a,F1),multiply(E1::'a,F1))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2879
\  (\\<forall>G1 I1 H1. equal(G1::'a,H1) --> equal(multiply(I1::'a,G1),multiply(I1::'a,H1))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2880
\  (\\<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)))))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2881
\  (\\<forall>L M N O_. equal(L::'a,M) --> equal(associator(L::'a,N,O_),associator(M::'a,N,O_))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2882
\  (\\<forall>P R Q S'. equal(P::'a,Q) --> equal(associator(R::'a,P,S'),associator(R::'a,Q,S'))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2883
\  (\\<forall>T' V W U. equal(T'::'a,U) --> equal(associator(V::'a,W,T'),associator(V::'a,W,U))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2884
\  (\\<forall>X Y. ~equal(multiply(multiply(Y::'a,X),Y),multiply(Y::'a,multiply(X::'a,Y)))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2885
\  (\\<forall>X Y Z. ~equal(associator(Y::'a,X,Z),additive_inverse(associator(X::'a,Y,Z)))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2886
\  (\\<forall>X Y Z. ~equal(associator(Z::'a,Y,X),additive_inverse(associator(X::'a,Y,Z)))) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2887
\  (~equal(multiply(multiply(cx::'a,multiply(cy::'a,cx)),cz),multiply(cx::'a,multiply(cy::'a,multiply(cx::'a,cz))))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2888
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2889
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2890
(*209 inferences so far.  Searching to depth 9.  1.2 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2891
val RNG038_2 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2892
 ("(\\<forall>X. sum(X::'a,additive_identity,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2893
\  (\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2894
\  (\\<forall>X Y. sum(X::'a,Y,add(X::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2895
\  (\\<forall>X. sum(X::'a,additive_inverse(X),additive_identity)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2896
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2897
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2898
\  (\\<forall>Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2899
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2900
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2901
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2902
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2903
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2904
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2905
\  (\\<forall>X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2906
\  (\\<forall>X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2907
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(additive_inverse(X),additive_inverse(Y))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2908
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(add(X::'a,W),add(Y::'a,W))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2909
\  (\\<forall>X Y W Z. equal(X::'a,Y) & sum(X::'a,W,Z) --> sum(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2910
\  (\\<forall>X W Y Z. equal(X::'a,Y) & sum(W::'a,X,Z) --> sum(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2911
\  (\\<forall>X W Z Y. equal(X::'a,Y) & sum(W::'a,Z,X) --> sum(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2912
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2913
\  (\\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2914
\  (\\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2915
\  (\\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2916
\  (\\<forall>X. product(additive_identity::'a,X,additive_identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2917
\  (\\<forall>X. product(X::'a,additive_identity,additive_identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2918
\  (\\<forall>X Y. equal(X::'a,additive_identity) --> product(X::'a,h(X::'a,Y),Y)) &        \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2919
\  (product(a::'a,b,additive_identity)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2920
\  (~equal(a::'a,additive_identity)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2921
\  (~equal(b::'a,additive_identity)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2922
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2923
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2924
(*2660 inferences so far.  Searching to depth 10.  7.0 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2925
val RNG040_2 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2926
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2927
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2928
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2929
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(additive_inverse(X),additive_inverse(Y))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2930
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(add(X::'a,W),add(Y::'a,W))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2931
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(add(W::'a,X),add(W::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2932
\  (\\<forall>X Y W Z. equal(X::'a,Y) & sum(X::'a,W,Z) --> sum(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2933
\  (\\<forall>X W Y Z. equal(X::'a,Y) & sum(W::'a,X,Z) --> sum(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2934
\  (\\<forall>X W Z Y. equal(X::'a,Y) & sum(W::'a,Z,X) --> sum(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2935
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2936
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2937
\  (\\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2938
\  (\\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2939
\  (\\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2940
\  (\\<forall>X. sum(additive_identity::'a,X,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2941
\  (\\<forall>X. sum(X::'a,additive_identity,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2942
\  (\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2943
\  (\\<forall>X Y. sum(X::'a,Y,add(X::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2944
\  (\\<forall>X. sum(additive_inverse(X),X,additive_identity)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2945
\  (\\<forall>X. sum(X::'a,additive_inverse(X),additive_identity)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2946
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2947
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2948
\  (\\<forall>Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2949
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2950
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2951
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2952
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2953
\  (\\<forall>X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2954
\  (\\<forall>X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2955
\  (\\<forall>A. product(A::'a,multiplicative_identity,A)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2956
\  (\\<forall>A. product(multiplicative_identity::'a,A,A)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2957
\  (\\<forall>A. product(A::'a,h(A),multiplicative_identity) | equal(A::'a,additive_identity)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2958
\  (\\<forall>A. product(h(A),A,multiplicative_identity) | equal(A::'a,additive_identity)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2959
\  (\\<forall>B A C. product(A::'a,B,C) --> product(B::'a,A,C)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2960
\  (\\<forall>A B. equal(A::'a,B) --> equal(h(A),h(B))) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2961
\  (sum(b::'a,c,d)) &       \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2962
\  (product(d::'a,a,additive_identity)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2963
\  (product(b::'a,a,l)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2964
\  (product(c::'a,a,n)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2965
\  (~sum(l::'a,n,additive_identity)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  2966
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2967
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2968
(*8991 inferences so far.  Searching to depth 9.  22.2 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  2969
val RNG041_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2970
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2971
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2972
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2973
\  (\\<forall>X. sum(additive_identity::'a,X,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2974
\  (\\<forall>X. sum(X::'a,additive_identity,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2975
\  (\\<forall>X Y. product(X::'a,Y,multiply(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2976
\  (\\<forall>X Y. sum(X::'a,Y,add(X::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2977
\  (\\<forall>X. sum(additive_inverse(X),X,additive_identity)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2978
\  (\\<forall>X. sum(X::'a,additive_inverse(X),additive_identity)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2979
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2980
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2981
\  (\\<forall>Y X Z. sum(X::'a,Y,Z) --> sum(Y::'a,X,Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2982
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2983
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2984
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2985
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2986
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2987
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2988
\  (\\<forall>X Y U V. sum(X::'a,Y,U) & sum(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2989
\  (\\<forall>X Y U V. product(X::'a,Y,U) & product(X::'a,Y,V) --> equal(U::'a,V)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2990
\  (\\<forall>X Y. equal(X::'a,Y) --> equal(additive_inverse(X),additive_inverse(Y))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2991
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(add(X::'a,W),add(Y::'a,W))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2992
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(add(W::'a,X),add(W::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2993
\  (\\<forall>X Y W Z. equal(X::'a,Y) & sum(X::'a,W,Z) --> sum(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2994
\  (\\<forall>X W Y Z. equal(X::'a,Y) & sum(W::'a,X,Z) --> sum(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2995
\  (\\<forall>X W Z Y. equal(X::'a,Y) & sum(W::'a,Z,X) --> sum(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2996
\  (\\<forall>X Y W. equal(X::'a,Y) --> equal(multiply(X::'a,W),multiply(Y::'a,W))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2997
\  (\\<forall>X W Y. equal(X::'a,Y) --> equal(multiply(W::'a,X),multiply(W::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2998
\  (\\<forall>X Y W Z. equal(X::'a,Y) & product(X::'a,W,Z) --> product(Y::'a,W,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  2999
\  (\\<forall>X W Y Z. equal(X::'a,Y) & product(W::'a,X,Z) --> product(W::'a,Y,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3000
\  (\\<forall>X W Z Y. equal(X::'a,Y) & product(W::'a,Z,X) --> product(W::'a,Z,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3001
\  (\\<forall>A B. equal(A::'a,B) --> equal(h(A),h(B))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3002
\  (\\<forall>A. product(additive_identity::'a,A,additive_identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3003
\  (\\<forall>A. product(A::'a,additive_identity,additive_identity)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3004
\  (\\<forall>A. product(A::'a,multiplicative_identity,A)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3005
\  (\\<forall>A. product(multiplicative_identity::'a,A,A)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3006
\  (\\<forall>A. product(A::'a,h(A),multiplicative_identity) | equal(A::'a,additive_identity)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3007
\  (\\<forall>A. product(h(A),A,multiplicative_identity) | equal(A::'a,additive_identity)) &       \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3008
\  (product(a::'a,b,additive_identity)) &   \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3009
\  (~equal(a::'a,additive_identity)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3010
\  (~equal(b::'a,additive_identity)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3011
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3012
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3013
(*101319 inferences so far.  Searching to depth 14.  76.0 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3014
val ROB010_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3015
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3016
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3017
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3018
\  (\\<forall>Y X. equal(add(X::'a,Y),add(Y::'a,X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3019
\  (\\<forall>X Y Z. equal(add(add(X::'a,Y),Z),add(X::'a,add(Y::'a,Z)))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3020
\  (\\<forall>Y X. equal(negate(add(negate(add(X::'a,Y)),negate(add(X::'a,negate(Y))))),X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3021
\  (\\<forall>A B C. equal(A::'a,B) --> equal(add(A::'a,C),add(B::'a,C))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3022
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(add(F'::'a,D),add(F'::'a,E))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3023
\  (\\<forall>G H. equal(G::'a,H) --> equal(negate(G),negate(H))) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3024
\  (equal(negate(add(a::'a,negate(b))),c)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3025
\  (~equal(negate(add(c::'a,negate(add(b::'a,a)))),a)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3026
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3027
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3028
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3029
(*6933 inferences so far.  Searching to depth 12.  5.1 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3030
val ROB013_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3031
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3032
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3033
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3034
\  (\\<forall>Y X. equal(add(X::'a,Y),add(Y::'a,X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3035
\  (\\<forall>X Y Z. equal(add(add(X::'a,Y),Z),add(X::'a,add(Y::'a,Z)))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3036
\  (\\<forall>Y X. equal(negate(add(negate(add(X::'a,Y)),negate(add(X::'a,negate(Y))))),X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3037
\  (\\<forall>A B C. equal(A::'a,B) --> equal(add(A::'a,C),add(B::'a,C))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3038
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(add(F'::'a,D),add(F'::'a,E))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3039
\  (\\<forall>G H. equal(G::'a,H) --> equal(negate(G),negate(H))) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3040
\  (equal(negate(add(a::'a,b)),c)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3041
\  (~equal(negate(add(c::'a,negate(add(negate(b),a)))),a)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3042
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3043
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3044
(*6614 inferences so far.  Searching to depth 11.  20.4 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3045
val ROB016_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3046
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3047
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3048
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3049
\  (\\<forall>Y X. equal(add(X::'a,Y),add(Y::'a,X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3050
\  (\\<forall>X Y Z. equal(add(add(X::'a,Y),Z),add(X::'a,add(Y::'a,Z)))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3051
\  (\\<forall>Y X. equal(negate(add(negate(add(X::'a,Y)),negate(add(X::'a,negate(Y))))),X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3052
\  (\\<forall>A B C. equal(A::'a,B) --> equal(add(A::'a,C),add(B::'a,C))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3053
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(add(F'::'a,D),add(F'::'a,E))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3054
\  (\\<forall>G H. equal(G::'a,H) --> equal(negate(G),negate(H))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3055
\  (\\<forall>J K' L. equal(J::'a,K') --> equal(multiply(J::'a,L),multiply(K'::'a,L))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3056
\  (\\<forall>M O_ N. equal(M::'a,N) --> equal(multiply(O_::'a,M),multiply(O_::'a,N))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3057
\  (\\<forall>P Q. equal(P::'a,Q) --> equal(successor(P),successor(Q))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3058
\  (\\<forall>R S'. equal(R::'a,S') & positive_integer(R) --> positive_integer(S')) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3059
\  (\\<forall>X. equal(multiply(one::'a,X),X)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3060
\  (\\<forall>V X. positive_integer(X) --> equal(multiply(successor(V),X),add(X::'a,multiply(V::'a,X)))) &      \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3061
\  (positive_integer(one)) &    \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3062
\  (\\<forall>X. positive_integer(X) --> positive_integer(successor(X))) &      \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3063
\  (equal(negate(add(d::'a,e)),negate(e))) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3064
\  (positive_integer(k)) &      \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3065
\  (\\<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))) &       \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3066
\  (~equal(negate(add(e::'a,multiply(k::'a,add(d::'a,negate(add(d::'a,negate(e))))))),negate(e))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3067
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3068
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3069
(*14077 inferences so far.  Searching to depth 11.  32.8 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3070
val ROB021_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3071
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3072
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3073
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3074
\  (\\<forall>Y X. equal(add(X::'a,Y),add(Y::'a,X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3075
\  (\\<forall>X Y Z. equal(add(add(X::'a,Y),Z),add(X::'a,add(Y::'a,Z)))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3076
\  (\\<forall>Y X. equal(negate(add(negate(add(X::'a,Y)),negate(add(X::'a,negate(Y))))),X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3077
\  (\\<forall>A B C. equal(A::'a,B) --> equal(add(A::'a,C),add(B::'a,C))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3078
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(add(F'::'a,D),add(F'::'a,E))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3079
\  (\\<forall>G H. equal(G::'a,H) --> equal(negate(G),negate(H))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3080
\  (\\<forall>X Y. equal(negate(X),negate(Y)) --> equal(X::'a,Y)) & \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3081
\  (~equal(add(negate(add(a::'a,negate(b))),negate(add(negate(a),negate(b)))),b)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3082
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3083
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3084
(*35532 inferences so far.  Searching to depth 19.  54.3 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3085
val SET005_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3086
 ("(\\<forall>Subset Element Superset. member(Element::'a,Subset) & subset(Subset::'a,Superset) --> member(Element::'a,Superset)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3087
\  (\\<forall>Superset Subset. subset(Subset::'a,Superset) | member(member_of_1_not_of_2(Subset::'a,Superset),Subset)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3088
\  (\\<forall>Subset Superset. member(member_of_1_not_of_2(Subset::'a,Superset),Superset) --> subset(Subset::'a,Superset)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3089
\  (\\<forall>Subset Superset. equal_sets(Subset::'a,Superset) --> subset(Subset::'a,Superset)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3090
\  (\\<forall>Subset Superset. equal_sets(Superset::'a,Subset) --> subset(Subset::'a,Superset)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3091
\  (\\<forall>Set2 Set1. subset(Set1::'a,Set2) & subset(Set2::'a,Set1) --> equal_sets(Set2::'a,Set1)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3092
\  (\\<forall>Set2 Intersection Element Set1. intersection(Set1::'a,Set2,Intersection) & member(Element::'a,Intersection) --> member(Element::'a,Set1)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3093
\  (\\<forall>Set1 Intersection Element Set2. intersection(Set1::'a,Set2,Intersection) & member(Element::'a,Intersection) --> member(Element::'a,Set2)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3094
\  (\\<forall>Set2 Set1 Element Intersection. intersection(Set1::'a,Set2,Intersection) & member(Element::'a,Set2) & member(Element::'a,Set1) --> member(Element::'a,Intersection)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3095
\  (\\<forall>Set2 Intersection Set1. member(h(Set1::'a,Set2,Intersection),Intersection) | intersection(Set1::'a,Set2,Intersection) | member(h(Set1::'a,Set2,Intersection),Set1)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3096
\  (\\<forall>Set1 Intersection Set2. member(h(Set1::'a,Set2,Intersection),Intersection) | intersection(Set1::'a,Set2,Intersection) | member(h(Set1::'a,Set2,Intersection),Set2)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3097
\  (\\<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)) &      \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3098
\  (intersection(a::'a,b,aIb)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3099
\  (intersection(b::'a,c,bIc)) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3100
\  (intersection(a::'a,bIc,aIbIc)) &        \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3101
\  (~intersection(aIb::'a,c,aIbIc)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3102
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3103
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3104
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3105
(*6450 inferences so far.  Searching to depth 14.  4.2 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3106
val SET009_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3107
 ("(\\<forall>Subset Element Superset. member(Element::'a,Subset) & subset(Subset::'a,Superset) --> member(Element::'a,Superset)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3108
\  (\\<forall>Superset Subset. subset(Subset::'a,Superset) | member(member_of_1_not_of_2(Subset::'a,Superset),Subset)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3109
\  (\\<forall>Subset Superset. member(member_of_1_not_of_2(Subset::'a,Superset),Superset) --> subset(Subset::'a,Superset)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3110
\  (\\<forall>Subset Superset. equal_sets(Subset::'a,Superset) --> subset(Subset::'a,Superset)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3111
\  (\\<forall>Subset Superset. equal_sets(Superset::'a,Subset) --> subset(Subset::'a,Superset)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3112
\  (\\<forall>Set2 Set1. subset(Set1::'a,Set2) & subset(Set2::'a,Set1) --> equal_sets(Set2::'a,Set1)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3113
\  (\\<forall>Set2 Difference Element Set1. difference(Set1::'a,Set2,Difference) & member(Element::'a,Difference) --> member(Element::'a,Set1)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3114
\  (\\<forall>Element A_set Set1 Set2. ~(member(Element::'a,Set1) & member(Element::'a,Set2) & difference(A_set::'a,Set1,Set2))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3115
\  (\\<forall>Set1 Difference Element Set2. member(Element::'a,Set1) & difference(Set1::'a,Set2,Difference) --> member(Element::'a,Difference) | member(Element::'a,Set2)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3116
\  (\\<forall>Set1 Set2 Difference. difference(Set1::'a,Set2,Difference) | member(k(Set1::'a,Set2,Difference),Set1) | member(k(Set1::'a,Set2,Difference),Difference)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3117
\  (\\<forall>Set1 Set2 Difference. member(k(Set1::'a,Set2,Difference),Set2) --> member(k(Set1::'a,Set2,Difference),Difference) | difference(Set1::'a,Set2,Difference)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3118
\  (\\<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)) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3119
\  (subset(d::'a,a)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3120
\  (difference(b::'a,a,bDa)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3121
\  (difference(b::'a,d,bDd)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3122
\  (~subset(bDa::'a,bDd)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3123
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3124
14220
4dc132902672 Merging of ex/cla.ML and ex/mesontest.ML to ex/Classical.thy
paulson
parents: 14183
diff changeset
  3125
(*34726 inferences so far.  Searching to depth 6.  2420 secs: 40 mins! BIG*)
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3126
val SET025_4 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3127
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3128
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3129
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3130
\  (\\<forall>Y X. member(X::'a,Y) --> little_set(X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3131
\  (\\<forall>X Y. little_set(f1(X::'a,Y)) | equal(X::'a,Y)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3132
\  (\\<forall>X Y. member(f1(X::'a,Y),X) | member(f1(X::'a,Y),Y) | equal(X::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3133
\  (\\<forall>X Y. member(f1(X::'a,Y),X) & member(f1(X::'a,Y),Y) --> equal(X::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3134
\  (\\<forall>X U Y. member(U::'a,non_ordered_pair(X::'a,Y)) --> equal(U::'a,X) | equal(U::'a,Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3135
\  (\\<forall>Y U X. little_set(U) & equal(U::'a,X) --> member(U::'a,non_ordered_pair(X::'a,Y))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3136
\  (\\<forall>X U Y. little_set(U) & equal(U::'a,Y) --> member(U::'a,non_ordered_pair(X::'a,Y))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3137
\  (\\<forall>X Y. little_set(non_ordered_pair(X::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3138
\  (\\<forall>X. equal(singleton_set(X),non_ordered_pair(X::'a,X))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3139
\  (\\<forall>X Y. equal(ordered_pair(X::'a,Y),non_ordered_pair(singleton_set(X),non_ordered_pair(X::'a,Y)))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3140
\  (\\<forall>X. ordered_pair_predicate(X) --> little_set(f2(X))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3141
\  (\\<forall>X. ordered_pair_predicate(X) --> little_set(f3(X))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3142
\  (\\<forall>X. ordered_pair_predicate(X) --> equal(X::'a,ordered_pair(f2(X),f3(X)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3143
\  (\\<forall>X Y Z. little_set(Y) & little_set(Z) & equal(X::'a,ordered_pair(Y::'a,Z)) --> ordered_pair_predicate(X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3144
\  (\\<forall>Z X. member(Z::'a,first(X)) --> little_set(f4(Z::'a,X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3145
\  (\\<forall>Z X. member(Z::'a,first(X)) --> little_set(f5(Z::'a,X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3146
\  (\\<forall>Z X. member(Z::'a,first(X)) --> equal(X::'a,ordered_pair(f4(Z::'a,X),f5(Z::'a,X)))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3147
\  (\\<forall>Z X. member(Z::'a,first(X)) --> member(Z::'a,f4(Z::'a,X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3148
\  (\\<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))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3149
\  (\\<forall>Z X. member(Z::'a,second(X)) --> little_set(f6(Z::'a,X))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3150
\  (\\<forall>Z X. member(Z::'a,second(X)) --> little_set(f7(Z::'a,X))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3151
\  (\\<forall>Z X. member(Z::'a,second(X)) --> equal(X::'a,ordered_pair(f6(Z::'a,X),f7(Z::'a,X)))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3152
\  (\\<forall>Z X. member(Z::'a,second(X)) --> member(Z::'a,f7(Z::'a,X))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3153
\  (\\<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))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3154
\  (\\<forall>Z. member(Z::'a,estin) --> ordered_pair_predicate(Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3155
\  (\\<forall>Z. member(Z::'a,estin) --> member(first(Z),second(Z))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3156
\  (\\<forall>Z. little_set(Z) & ordered_pair_predicate(Z) & member(first(Z),second(Z)) --> member(Z::'a,estin)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3157
\  (\\<forall>Y Z X. member(Z::'a,intersection(X::'a,Y)) --> member(Z::'a,X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3158
\  (\\<forall>X Z Y. member(Z::'a,intersection(X::'a,Y)) --> member(Z::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3159
\  (\\<forall>X Z Y. member(Z::'a,X) & member(Z::'a,Y) --> member(Z::'a,intersection(X::'a,Y))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3160
\  (\\<forall>Z X. ~(member(Z::'a,complement(X)) & member(Z::'a,X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3161
\  (\\<forall>Z X. little_set(Z) --> member(Z::'a,complement(X)) | member(Z::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3162
\  (\\<forall>X Y. equal(union(X::'a,Y),complement(intersection(complement(X),complement(Y))))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3163
\  (\\<forall>Z X. member(Z::'a,domain_of(X)) --> ordered_pair_predicate(f8(Z::'a,X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3164
\  (\\<forall>Z X. member(Z::'a,domain_of(X)) --> member(f8(Z::'a,X),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3165
\  (\\<forall>Z X. member(Z::'a,domain_of(X)) --> equal(Z::'a,first(f8(Z::'a,X)))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3166
\  (\\<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))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3167
\  (\\<forall>X Y Z. member(Z::'a,cross_product(X::'a,Y)) --> ordered_pair_predicate(Z)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3168
\  (\\<forall>Y Z X. member(Z::'a,cross_product(X::'a,Y)) --> member(first(Z),X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3169
\  (\\<forall>X Z Y. member(Z::'a,cross_product(X::'a,Y)) --> member(second(Z),Y)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3170
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3171
\  (\\<forall>X Z. member(Z::'a,inv1 X) --> ordered_pair_predicate(Z)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3172
\  (\\<forall>Z X. member(Z::'a,inv1 X) --> member(ordered_pair(second(Z),first(Z)),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3173
\  (\\<forall>Z X. little_set(Z) & ordered_pair_predicate(Z) & member(ordered_pair(second(Z),first(Z)),X) --> member(Z::'a,inv1 X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3174
\  (\\<forall>Z X. member(Z::'a,rotate_right(X)) --> little_set(f9(Z::'a,X))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3175
\  (\\<forall>Z X. member(Z::'a,rotate_right(X)) --> little_set(f10(Z::'a,X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3176
\  (\\<forall>Z X. member(Z::'a,rotate_right(X)) --> little_set(f11(Z::'a,X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3177
\  (\\<forall>Z X. member(Z::'a,rotate_right(X)) --> equal(Z::'a,ordered_pair(f9(Z::'a,X),ordered_pair(f10(Z::'a,X),f11(Z::'a,X))))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3178
\  (\\<forall>Z X. member(Z::'a,rotate_right(X)) --> member(ordered_pair(f10(Z::'a,X),ordered_pair(f11(Z::'a,X),f9(Z::'a,X))),X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3179
\  (\\<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,rotate_right(X))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3180
\  (\\<forall>Z X. member(Z::'a,flip_range_of(X)) --> little_set(f12(Z::'a,X))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3181
\  (\\<forall>Z X. member(Z::'a,flip_range_of(X)) --> little_set(f13(Z::'a,X))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3182
\  (\\<forall>Z X. member(Z::'a,flip_range_of(X)) --> little_set(f14(Z::'a,X))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3183
\  (\\<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))))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3184
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3185
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3186
\  (\\<forall>X. equal(successor(X),union(X::'a,singleton_set(X)))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3187
\  (\\<forall>Z. ~member(Z::'a,empty_set)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3188
\  (\\<forall>Z. little_set(Z) --> member(Z::'a,universal_set)) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3189
\  (little_set(infinity)) &     \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3190
\  (member(empty_set::'a,infinity)) &       \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3191
\  (\\<forall>X. member(X::'a,infinity) --> member(successor(X),infinity)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3192
\  (\\<forall>Z X. member(Z::'a,sigma(X)) --> member(f16(Z::'a,X),X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3193
\  (\\<forall>Z X. member(Z::'a,sigma(X)) --> member(Z::'a,f16(Z::'a,X))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3194
\  (\\<forall>X Z Y. member(Y::'a,X) & member(Z::'a,Y) --> member(Z::'a,sigma(X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3195
\  (\\<forall>U. little_set(U) --> little_set(sigma(U))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3196
\  (\\<forall>X U Y. subset(X::'a,Y) & member(U::'a,X) --> member(U::'a,Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3197
\  (\\<forall>Y X. subset(X::'a,Y) | member(f17(X::'a,Y),X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3198
\  (\\<forall>X Y. member(f17(X::'a,Y),Y) --> subset(X::'a,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3199
\  (\\<forall>X Y. proper_subset(X::'a,Y) --> subset(X::'a,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3200
\  (\\<forall>X Y. ~(proper_subset(X::'a,Y) & equal(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3201
\  (\\<forall>X Y. subset(X::'a,Y) --> proper_subset(X::'a,Y) | equal(X::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3202
\  (\\<forall>Z X. member(Z::'a,powerset(X)) --> subset(Z::'a,X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3203
\  (\\<forall>Z X. little_set(Z) & subset(Z::'a,X) --> member(Z::'a,powerset(X))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3204
\  (\\<forall>U. little_set(U) --> little_set(powerset(U))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3205
\  (\\<forall>Z X. relation(Z) & member(X::'a,Z) --> ordered_pair_predicate(X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3206
\  (\\<forall>Z. relation(Z) | member(f18(Z),Z)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3207
\  (\\<forall>Z. ordered_pair_predicate(f18(Z)) --> relation(Z)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3208
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3209
\  (\\<forall>X. single_valued_set(X) | little_set(f19(X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3210
\  (\\<forall>X. single_valued_set(X) | little_set(f20(X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3211
\  (\\<forall>X. single_valued_set(X) | little_set(f21(X))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3212
\  (\\<forall>X. single_valued_set(X) | member(ordered_pair(f19(X),f20(X)),X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3213
\  (\\<forall>X. single_valued_set(X) | member(ordered_pair(f19(X),f21(X)),X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3214
\  (\\<forall>X. equal(f20(X),f21(X)) --> single_valued_set(X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3215
\  (\\<forall>Xf. function(Xf) --> relation(Xf)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3216
\  (\\<forall>Xf. function(Xf) --> single_valued_set(Xf)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3217
\  (\\<forall>Xf. relation(Xf) & single_valued_set(Xf) --> function(Xf)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3218
\  (\\<forall>Z X Xf. member(Z::'a,image_(X::'a,Xf)) --> ordered_pair_predicate(f22(Z::'a,X,Xf))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3219
\  (\\<forall>Z X Xf. member(Z::'a,image_(X::'a,Xf)) --> member(f22(Z::'a,X,Xf),Xf)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3220
\  (\\<forall>Z Xf X. member(Z::'a,image_(X::'a,Xf)) --> member(first(f22(Z::'a,X,Xf)),X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3221
\  (\\<forall>X Xf Z. member(Z::'a,image_(X::'a,Xf)) --> equal(second(f22(Z::'a,X,Xf)),Z)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3222
\  (\\<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))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3223
\  (\\<forall>X Xf. little_set(X) & function(Xf) --> little_set(image_(X::'a,Xf))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3224
\  (\\<forall>X U Y. ~(disjoint(X::'a,Y) & member(U::'a,X) & member(U::'a,Y))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3225
\  (\\<forall>Y X. disjoint(X::'a,Y) | member(f23(X::'a,Y),X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3226
\  (\\<forall>X Y. disjoint(X::'a,Y) | member(f23(X::'a,Y),Y)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3227
\  (\\<forall>X. equal(X::'a,empty_set) | member(f24(X),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3228
\  (\\<forall>X. equal(X::'a,empty_set) | disjoint(f24(X),X)) &    \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3229
\  (function(f25)) &    \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3230
\  (\\<forall>X. little_set(X) --> equal(X::'a,empty_set) | member(f26(X),X)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3231
\  (\\<forall>X. little_set(X) --> equal(X::'a,empty_set) | member(ordered_pair(X::'a,f26(X)),f25)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3232
\  (\\<forall>Z X. member(Z::'a,range_of(X)) --> ordered_pair_predicate(f27(Z::'a,X))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3233
\  (\\<forall>Z X. member(Z::'a,range_of(X)) --> member(f27(Z::'a,X),X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3234
\  (\\<forall>Z X. member(Z::'a,range_of(X)) --> equal(Z::'a,second(f27(Z::'a,X)))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3235
\  (\\<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))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3236
\  (\\<forall>Z. member(Z::'a,identity_relation) --> ordered_pair_predicate(Z)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3237
\  (\\<forall>Z. member(Z::'a,identity_relation) --> equal(first(Z),second(Z))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3238
\  (\\<forall>Z. little_set(Z) & ordered_pair_predicate(Z) & equal(first(Z),second(Z)) --> member(Z::'a,identity_relation)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3239
\  (\\<forall>X Y. equal(restrct(X::'a,Y),intersection(X::'a,cross_product(Y::'a,universal_set)))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3240
\  (\\<forall>Xf. one_to_one_function(Xf) --> function(Xf)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3241
\  (\\<forall>Xf. one_to_one_function(Xf) --> function(inv1 Xf)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3242
\  (\\<forall>Xf. function(Xf) & function(inv1 Xf) --> one_to_one_function(Xf)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3243
\  (\\<forall>Z Xf Y. member(Z::'a,apply(Xf::'a,Y)) --> ordered_pair_predicate(f28(Z::'a,Xf,Y))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3244
\  (\\<forall>Z Y Xf. member(Z::'a,apply(Xf::'a,Y)) --> member(f28(Z::'a,Xf,Y),Xf)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3245
\  (\\<forall>Z Xf Y. member(Z::'a,apply(Xf::'a,Y)) --> equal(first(f28(Z::'a,Xf,Y)),Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3246
\  (\\<forall>Z Xf Y. member(Z::'a,apply(Xf::'a,Y)) --> member(Z::'a,second(f28(Z::'a,Xf,Y)))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3247
\  (\\<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))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3248
\  (\\<forall>Xf X Y. equal(apply_to_two_arguments(Xf::'a,X,Y),apply(Xf::'a,ordered_pair(X::'a,Y)))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3249
\  (\\<forall>X Y Xf. maps(Xf::'a,X,Y) --> function(Xf)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3250
\  (\\<forall>Y Xf X. maps(Xf::'a,X,Y) --> equal(domain_of(Xf),X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3251
\  (\\<forall>X Xf Y. maps(Xf::'a,X,Y) --> subset(range_of(Xf),Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3252
\  (\\<forall>X Xf Y. function(Xf) & equal(domain_of(Xf),X) & subset(range_of(Xf),Y) --> maps(Xf::'a,X,Y)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3253
\  (\\<forall>Xf Xs. closed(Xs::'a,Xf) --> little_set(Xs)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3254
\  (\\<forall>Xs Xf. closed(Xs::'a,Xf) --> little_set(Xf)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3255
\  (\\<forall>Xf Xs. closed(Xs::'a,Xf) --> maps(Xf::'a,cross_product(Xs::'a,Xs),Xs)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3256
\  (\\<forall>Xf Xs. little_set(Xs) & little_set(Xf) & maps(Xf::'a,cross_product(Xs::'a,Xs),Xs) --> closed(Xs::'a,Xf)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3257
\  (\\<forall>Z Xf Xg. member(Z::'a,composition(Xf::'a,Xg)) --> little_set(f29(Z::'a,Xf,Xg))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3258
\  (\\<forall>Z Xf Xg. member(Z::'a,composition(Xf::'a,Xg)) --> little_set(f30(Z::'a,Xf,Xg))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3259
\  (\\<forall>Z Xf Xg. member(Z::'a,composition(Xf::'a,Xg)) --> little_set(f31(Z::'a,Xf,Xg))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3260
\  (\\<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)))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3261
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3262
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3263
\  (\\<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))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3264
\  (\\<forall>Xh Xs2 Xf2 Xs1 Xf1. homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2) --> closed(Xs1::'a,Xf1)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3265
\  (\\<forall>Xh Xs1 Xf1 Xs2 Xf2. homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2) --> closed(Xs2::'a,Xf2)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3266
\  (\\<forall>Xf1 Xf2 Xh Xs1 Xs2. homomorphism(Xh::'a,Xs1,Xf1,Xs2,Xf2) --> maps(Xh::'a,Xs1,Xs2)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3267
\  (\\<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)))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3268
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3269
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3270
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3271
\  (\\<forall>A B C. equal(A::'a,B) --> equal(f1(A::'a,C),f1(B::'a,C))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3272
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(f1(F'::'a,D),f1(F'::'a,E))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3273
\  (\\<forall>A2 B2. equal(A2::'a,B2) --> equal(f2(A2),f2(B2))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3274
\  (\\<forall>G4 H4. equal(G4::'a,H4) --> equal(f3(G4),f3(H4))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3275
\  (\\<forall>O7 P7 Q7. equal(O7::'a,P7) --> equal(f4(O7::'a,Q7),f4(P7::'a,Q7))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3276
\  (\\<forall>R7 T7 S7. equal(R7::'a,S7) --> equal(f4(T7::'a,R7),f4(T7::'a,S7))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3277
\  (\\<forall>U7 V7 W7. equal(U7::'a,V7) --> equal(f5(U7::'a,W7),f5(V7::'a,W7))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3278
\  (\\<forall>X7 Z7 Y7. equal(X7::'a,Y7) --> equal(f5(Z7::'a,X7),f5(Z7::'a,Y7))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3279
\  (\\<forall>A8 B8 C8. equal(A8::'a,B8) --> equal(f6(A8::'a,C8),f6(B8::'a,C8))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3280
\  (\\<forall>D8 F8 E8. equal(D8::'a,E8) --> equal(f6(F8::'a,D8),f6(F8::'a,E8))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3281
\  (\\<forall>G8 H8 I8. equal(G8::'a,H8) --> equal(f7(G8::'a,I8),f7(H8::'a,I8))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3282
\  (\\<forall>J8 L8 K8. equal(J8::'a,K8) --> equal(f7(L8::'a,J8),f7(L8::'a,K8))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3283
\  (\\<forall>M8 N8 O8. equal(M8::'a,N8) --> equal(f8(M8::'a,O8),f8(N8::'a,O8))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3284
\  (\\<forall>P8 R8 Q8. equal(P8::'a,Q8) --> equal(f8(R8::'a,P8),f8(R8::'a,Q8))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3285
\  (\\<forall>S8 T8 U8. equal(S8::'a,T8) --> equal(f9(S8::'a,U8),f9(T8::'a,U8))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3286
\  (\\<forall>V8 X8 W8. equal(V8::'a,W8) --> equal(f9(X8::'a,V8),f9(X8::'a,W8))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3287
\  (\\<forall>G H I'. equal(G::'a,H) --> equal(f10(G::'a,I'),f10(H::'a,I'))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3288
\  (\\<forall>J L K'. equal(J::'a,K') --> equal(f10(L::'a,J),f10(L::'a,K'))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3289
\  (\\<forall>M N O_. equal(M::'a,N) --> equal(f11(M::'a,O_),f11(N::'a,O_))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3290
\  (\\<forall>P R Q. equal(P::'a,Q) --> equal(f11(R::'a,P),f11(R::'a,Q))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3291
\  (\\<forall>S' T' U. equal(S'::'a,T') --> equal(f12(S'::'a,U),f12(T'::'a,U))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3292
\  (\\<forall>V X W. equal(V::'a,W) --> equal(f12(X::'a,V),f12(X::'a,W))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3293
\  (\\<forall>Y Z A1. equal(Y::'a,Z) --> equal(f13(Y::'a,A1),f13(Z::'a,A1))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3294
\  (\\<forall>B1 D1 C1. equal(B1::'a,C1) --> equal(f13(D1::'a,B1),f13(D1::'a,C1))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3295
\  (\\<forall>E1 F1 G1. equal(E1::'a,F1) --> equal(f14(E1::'a,G1),f14(F1::'a,G1))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3296
\  (\\<forall>H1 J1 I1. equal(H1::'a,I1) --> equal(f14(J1::'a,H1),f14(J1::'a,I1))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3297
\  (\\<forall>K1 L1 M1. equal(K1::'a,L1) --> equal(f16(K1::'a,M1),f16(L1::'a,M1))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3298
\  (\\<forall>N1 P1 O1. equal(N1::'a,O1) --> equal(f16(P1::'a,N1),f16(P1::'a,O1))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3299
\  (\\<forall>Q1 R1 S1. equal(Q1::'a,R1) --> equal(f17(Q1::'a,S1),f17(R1::'a,S1))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3300
\  (\\<forall>T1 V1 U1. equal(T1::'a,U1) --> equal(f17(V1::'a,T1),f17(V1::'a,U1))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3301
\  (\\<forall>W1 X1. equal(W1::'a,X1) --> equal(f18(W1),f18(X1))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3302
\  (\\<forall>Y1 Z1. equal(Y1::'a,Z1) --> equal(f19(Y1),f19(Z1))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3303
\  (\\<forall>C2 D2. equal(C2::'a,D2) --> equal(f20(C2),f20(D2))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3304
\  (\\<forall>E2 F2. equal(E2::'a,F2) --> equal(f21(E2),f21(F2))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3305
\  (\\<forall>G2 H2 I2 J2. equal(G2::'a,H2) --> equal(f22(G2::'a,I2,J2),f22(H2::'a,I2,J2))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3306
\  (\\<forall>K2 M2 L2 N2. equal(K2::'a,L2) --> equal(f22(M2::'a,K2,N2),f22(M2::'a,L2,N2))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3307
\  (\\<forall>O2 Q2 R2 P2. equal(O2::'a,P2) --> equal(f22(Q2::'a,R2,O2),f22(Q2::'a,R2,P2))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3308
\  (\\<forall>S2 T2 U2. equal(S2::'a,T2) --> equal(f23(S2::'a,U2),f23(T2::'a,U2))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3309
\  (\\<forall>V2 X2 W2. equal(V2::'a,W2) --> equal(f23(X2::'a,V2),f23(X2::'a,W2))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3310
\  (\\<forall>Y2 Z2. equal(Y2::'a,Z2) --> equal(f24(Y2),f24(Z2))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3311
\  (\\<forall>A3 B3. equal(A3::'a,B3) --> equal(f26(A3),f26(B3))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3312
\  (\\<forall>C3 D3 E3. equal(C3::'a,D3) --> equal(f27(C3::'a,E3),f27(D3::'a,E3))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3313
\  (\\<forall>F3 H3 G3. equal(F3::'a,G3) --> equal(f27(H3::'a,F3),f27(H3::'a,G3))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3314
\  (\\<forall>I3 J3 K3 L3. equal(I3::'a,J3) --> equal(f28(I3::'a,K3,L3),f28(J3::'a,K3,L3))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3315
\  (\\<forall>M3 O3 N3 P3. equal(M3::'a,N3) --> equal(f28(O3::'a,M3,P3),f28(O3::'a,N3,P3))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3316
\  (\\<forall>Q3 S3 T3 R3. equal(Q3::'a,R3) --> equal(f28(S3::'a,T3,Q3),f28(S3::'a,T3,R3))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3317
\  (\\<forall>U3 V3 W3 X3. equal(U3::'a,V3) --> equal(f29(U3::'a,W3,X3),f29(V3::'a,W3,X3))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3318
\  (\\<forall>Y3 A4 Z3 B4. equal(Y3::'a,Z3) --> equal(f29(A4::'a,Y3,B4),f29(A4::'a,Z3,B4))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3319
\  (\\<forall>C4 E4 F4 D4. equal(C4::'a,D4) --> equal(f29(E4::'a,F4,C4),f29(E4::'a,F4,D4))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3320
\  (\\<forall>I4 J4 K4 L4. equal(I4::'a,J4) --> equal(f30(I4::'a,K4,L4),f30(J4::'a,K4,L4))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3321
\  (\\<forall>M4 O4 N4 P4. equal(M4::'a,N4) --> equal(f30(O4::'a,M4,P4),f30(O4::'a,N4,P4))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3322
\  (\\<forall>Q4 S4 T4 R4. equal(Q4::'a,R4) --> equal(f30(S4::'a,T4,Q4),f30(S4::'a,T4,R4))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3323
\  (\\<forall>U4 V4 W4 X4. equal(U4::'a,V4) --> equal(f31(U4::'a,W4,X4),f31(V4::'a,W4,X4))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3324
\  (\\<forall>Y4 A5 Z4 B5. equal(Y4::'a,Z4) --> equal(f31(A5::'a,Y4,B5),f31(A5::'a,Z4,B5))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3325
\  (\\<forall>C5 E5 F5 D5. equal(C5::'a,D5) --> equal(f31(E5::'a,F5,C5),f31(E5::'a,F5,D5))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3326
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3327
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3328
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3329
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3330
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3331
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3332
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3333
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3334
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3335
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3336
\  (\\<forall>A B C. equal(A::'a,B) --> equal(apply(A::'a,C),apply(B::'a,C))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3337
\  (\\<forall>D F' E. equal(D::'a,E) --> equal(apply(F'::'a,D),apply(F'::'a,E))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3338
\  (\\<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))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3339
\  (\\<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))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3340
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3341
\  (\\<forall>S' T'. equal(S'::'a,T') --> equal(complement(S'),complement(T'))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3342
\  (\\<forall>U V W. equal(U::'a,V) --> equal(composition(U::'a,W),composition(V::'a,W))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3343
\  (\\<forall>X Z Y. equal(X::'a,Y) --> equal(composition(Z::'a,X),composition(Z::'a,Y))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3344
\  (\\<forall>A1 B1. equal(A1::'a,B1) --> equal(inv1 A1,inv1 B1)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3345
\  (\\<forall>C1 D1 E1. equal(C1::'a,D1) --> equal(cross_product(C1::'a,E1),cross_product(D1::'a,E1))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3346
\  (\\<forall>F1 H1 G1. equal(F1::'a,G1) --> equal(cross_product(H1::'a,F1),cross_product(H1::'a,G1))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3347
\  (\\<forall>I1 J1. equal(I1::'a,J1) --> equal(domain_of(I1),domain_of(J1))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3348
\  (\\<forall>I10 J10. equal(I10::'a,J10) --> equal(first(I10),first(J10))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3349
\  (\\<forall>Q10 R10. equal(Q10::'a,R10) --> equal(flip_range_of(Q10),flip_range_of(R10))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3350
\  (\\<forall>S10 T10 U10. equal(S10::'a,T10) --> equal(image_(S10::'a,U10),image_(T10::'a,U10))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3351
\  (\\<forall>V10 X10 W10. equal(V10::'a,W10) --> equal(image_(X10::'a,V10),image_(X10::'a,W10))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3352
\  (\\<forall>Y10 Z10 A11. equal(Y10::'a,Z10) --> equal(intersection(Y10::'a,A11),intersection(Z10::'a,A11))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3353
\  (\\<forall>B11 D11 C11. equal(B11::'a,C11) --> equal(intersection(D11::'a,B11),intersection(D11::'a,C11))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3354
\  (\\<forall>E11 F11 G11. equal(E11::'a,F11) --> equal(non_ordered_pair(E11::'a,G11),non_ordered_pair(F11::'a,G11))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3355
\  (\\<forall>H11 J11 I11. equal(H11::'a,I11) --> equal(non_ordered_pair(J11::'a,H11),non_ordered_pair(J11::'a,I11))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3356
\  (\\<forall>K11 L11 M11. equal(K11::'a,L11) --> equal(ordered_pair(K11::'a,M11),ordered_pair(L11::'a,M11))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3357
\  (\\<forall>N11 P11 O11. equal(N11::'a,O11) --> equal(ordered_pair(P11::'a,N11),ordered_pair(P11::'a,O11))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3358
\  (\\<forall>Q11 R11. equal(Q11::'a,R11) --> equal(powerset(Q11),powerset(R11))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3359
\  (\\<forall>S11 T11. equal(S11::'a,T11) --> equal(range_of(S11),range_of(T11))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3360
\  (\\<forall>U11 V11 W11. equal(U11::'a,V11) --> equal(restrct(U11::'a,W11),restrct(V11::'a,W11))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3361
\  (\\<forall>X11 Z11 Y11. equal(X11::'a,Y11) --> equal(restrct(Z11::'a,X11),restrct(Z11::'a,Y11))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3362
\  (\\<forall>A12 B12. equal(A12::'a,B12) --> equal(rotate_right(A12),rotate_right(B12))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3363
\  (\\<forall>C12 D12. equal(C12::'a,D12) --> equal(second(C12),second(D12))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3364
\  (\\<forall>K12 L12. equal(K12::'a,L12) --> equal(sigma(K12),sigma(L12))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3365
\  (\\<forall>M12 N12. equal(M12::'a,N12) --> equal(singleton_set(M12),singleton_set(N12))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3366
\  (\\<forall>O12 P12. equal(O12::'a,P12) --> equal(successor(O12),successor(P12))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3367
\  (\\<forall>Q12 R12 S12. equal(Q12::'a,R12) --> equal(union(Q12::'a,S12),union(R12::'a,S12))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3368
\  (\\<forall>T12 V12 U12. equal(T12::'a,U12) --> equal(union(V12::'a,T12),union(V12::'a,U12))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3369
\  (\\<forall>W12 X12 Y12. equal(W12::'a,X12) & closed(W12::'a,Y12) --> closed(X12::'a,Y12)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3370
\  (\\<forall>Z12 B13 A13. equal(Z12::'a,A13) & closed(B13::'a,Z12) --> closed(B13::'a,A13)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3371
\  (\\<forall>C13 D13 E13. equal(C13::'a,D13) & disjoint(C13::'a,E13) --> disjoint(D13::'a,E13)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3372
\  (\\<forall>F13 H13 G13. equal(F13::'a,G13) & disjoint(H13::'a,F13) --> disjoint(H13::'a,G13)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3373
\  (\\<forall>I13 J13. equal(I13::'a,J13) & function(I13) --> function(J13)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3374
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3375
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3376
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3377
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3378
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3379
\  (\\<forall>O14 P14. equal(O14::'a,P14) & little_set(O14) --> little_set(P14)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3380
\  (\\<forall>Q14 R14 S14 T14. equal(Q14::'a,R14) & maps(Q14::'a,S14,T14) --> maps(R14::'a,S14,T14)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3381
\  (\\<forall>U14 W14 V14 X14. equal(U14::'a,V14) & maps(W14::'a,U14,X14) --> maps(W14::'a,V14,X14)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3382
\  (\\<forall>Y14 A15 B15 Z14. equal(Y14::'a,Z14) & maps(A15::'a,B15,Y14) --> maps(A15::'a,B15,Z14)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3383
\  (\\<forall>C15 D15 E15. equal(C15::'a,D15) & member(C15::'a,E15) --> member(D15::'a,E15)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3384
\  (\\<forall>F15 H15 G15. equal(F15::'a,G15) & member(H15::'a,F15) --> member(H15::'a,G15)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3385
\  (\\<forall>I15 J15. equal(I15::'a,J15) & one_to_one_function(I15) --> one_to_one_function(J15)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3386
\  (\\<forall>K15 L15. equal(K15::'a,L15) & ordered_pair_predicate(K15) --> ordered_pair_predicate(L15)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3387
\  (\\<forall>M15 N15 O15. equal(M15::'a,N15) & proper_subset(M15::'a,O15) --> proper_subset(N15::'a,O15)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3388
\  (\\<forall>P15 R15 Q15. equal(P15::'a,Q15) & proper_subset(R15::'a,P15) --> proper_subset(R15::'a,Q15)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3389
\  (\\<forall>S15 T15. equal(S15::'a,T15) & relation(S15) --> relation(T15)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3390
\  (\\<forall>U15 V15. equal(U15::'a,V15) & single_valued_set(U15) --> single_valued_set(V15)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3391
\  (\\<forall>W15 X15 Y15. equal(W15::'a,X15) & subset(W15::'a,Y15) --> subset(X15::'a,Y15)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3392
\  (\\<forall>Z15 B16 A16. equal(Z15::'a,A16) & subset(B16::'a,Z15) --> subset(B16::'a,A16)) &      \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3393
\  (~little_set(ordered_pair(a::'a,b))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3394
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3395
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3396
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3397
(*13 inferences so far.  Searching to depth 8.  0 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3398
val SET046_5 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3399
 ("(\\<forall>Y X. ~(element(X::'a,a) & element(X::'a,Y) & element(Y::'a,X))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3400
\  (\\<forall>X. element(X::'a,f(X)) | element(X::'a,a)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3401
\  (\\<forall>X. element(f(X),X) | element(X::'a,a)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3402
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3403
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3404
(*33 inferences so far.  Searching to depth 9.  0.2 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3405
val SET047_5 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3406
 ("(\\<forall>X Z Y. set_equal(X::'a,Y) & element(Z::'a,X) --> element(Z::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3407
\  (\\<forall>Y Z X. set_equal(X::'a,Y) & element(Z::'a,Y) --> element(Z::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3408
\  (\\<forall>X Y. element(f(X::'a,Y),X) | element(f(X::'a,Y),Y) | set_equal(X::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3409
\  (\\<forall>X Y. element(f(X::'a,Y),Y) & element(f(X::'a,Y),X) --> set_equal(X::'a,Y)) &  \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3410
\  (set_equal(a::'a,b) | set_equal(b::'a,a)) & \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3411
\  (~(set_equal(b::'a,a) & set_equal(a::'a,b))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3412
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3413
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3414
(*311 inferences so far.  Searching to depth 12.  0.1 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3415
val SYN034_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3416
 ("(\\<forall>A. p(A::'a,a) | p(A::'a,f(A))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3417
\  (\\<forall>A. p(A::'a,a) | p(f(A),A)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3418
\  (\\<forall>A B. ~(p(A::'a,B) & p(B::'a,A) & p(B::'a,a))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3419
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3420
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3421
(*30 inferences so far.  Searching to depth 6.  0.2 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3422
val SYN071_1 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3423
 ("(\\<forall>X. equal(X::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3424
\  (\\<forall>Y X. equal(X::'a,Y) --> equal(Y::'a,X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3425
\  (\\<forall>Y X Z. equal(X::'a,Y) & equal(Y::'a,Z) --> equal(X::'a,Z)) &  \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3426
\  (equal(a::'a,b) | equal(c::'a,d)) & \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3427
\  (equal(a::'a,c) | equal(b::'a,d)) & \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3428
\  (~equal(a::'a,d)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3429
\  (~equal(b::'a,c)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3430
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3431
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3432
(****************SLOW
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3433
655670 inferences so far.  Searching to depth 44.  No proof after 10 minutes.
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3434
val SYN349_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3435
 ("(\\<forall>X Y. f(w(X),g(X::'a,Y)) --> f(X::'a,g(X::'a,Y))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3436
\  (\\<forall>X Y. f(X::'a,g(X::'a,Y)) --> f(w(X),g(X::'a,Y))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3437
\  (\\<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))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3438
\  (\\<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))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3439
\  (\\<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))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3440
\  (\\<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))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3441
\  (\\<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))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3442
\  (\\<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))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3443
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3444
\  (\\<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",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3445
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3446
****************)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3447
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3448
(*398 inferences so far.  Searching to depth 12.  0.4 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3449
val SYN352_1 = prove
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3450
 ("(f(a::'a,b)) &   \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3451
\  (\\<forall>X Y. f(X::'a,Y) --> f(b::'a,z(X::'a,Y)) | f(Y::'a,z(X::'a,Y))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3452
\  (\\<forall>X Y. f(X::'a,Y) | f(z(X::'a,Y),z(X::'a,Y))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3453
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3454
\  (\\<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))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3455
\  (\\<forall>X Y. ~(f(X::'a,Y) & f(X::'a,z(X::'a,Y)) & f(Y::'a,z(X::'a,Y)))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3456
\  (\\<forall>X Y. f(X::'a,Y) --> f(X::'a,z(X::'a,Y)) | f(Y::'a,z(X::'a,Y))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3457
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3458
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3459
(*5336 inferences so far.  Searching to depth 15.  5.3 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3460
val TOP001_2 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3461
 ("(\\<forall>Vf U. element_of_set(U::'a,union_of_members(Vf)) --> element_of_set(U::'a,f1(Vf::'a,U))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3462
\  (\\<forall>U Vf. element_of_set(U::'a,union_of_members(Vf)) --> element_of_collection(f1(Vf::'a,U),Vf)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3463
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3464
\  (\\<forall>Vf X. basis(X::'a,Vf) --> equal_sets(union_of_members(Vf),X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3465
\  (\\<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))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3466
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3467
\  (\\<forall>X. subset_sets(X::'a,X)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3468
\  (\\<forall>X U Y. subset_sets(X::'a,Y) & element_of_set(U::'a,X) --> element_of_set(U::'a,Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3469
\  (\\<forall>X Y. equal_sets(X::'a,Y) --> subset_sets(X::'a,Y)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3470
\  (\\<forall>Y X. subset_sets(X::'a,Y) | element_of_set(in_1st_set(X::'a,Y),X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3471
\  (\\<forall>X Y. element_of_set(in_1st_set(X::'a,Y),Y) --> subset_sets(X::'a,Y)) &    \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3472
\  (basis(cx::'a,f)) &      \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3473
\  (~subset_sets(union_of_members(top_of_basis(f)),cx)) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3474
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3475
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3476
(*0 inferences so far.  Searching to depth 0.  0 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3477
val TOP002_2 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3478
 ("(\\<forall>Vf U. element_of_collection(U::'a,top_of_basis(Vf)) | element_of_set(f11(Vf::'a,U),U)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3479
\  (\\<forall>X. ~element_of_set(X::'a,empty_set)) &        \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3480
\  (~element_of_collection(empty_set::'a,top_of_basis(f))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3481
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3482
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3483
(*0 inferences so far.  Searching to depth 0.  6.5 secs.  BIG*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3484
val TOP004_1 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3485
 ("(\\<forall>Vf U. element_of_set(U::'a,union_of_members(Vf)) --> element_of_set(U::'a,f1(Vf::'a,U))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3486
\  (\\<forall>U Vf. element_of_set(U::'a,union_of_members(Vf)) --> element_of_collection(f1(Vf::'a,U),Vf)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3487
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3488
\  (\\<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)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3489
\  (\\<forall>U Vf. element_of_set(U::'a,intersection_of_members(Vf)) | element_of_collection(f2(Vf::'a,U),Vf)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3490
\  (\\<forall>Vf U. element_of_set(U::'a,f2(Vf::'a,U)) --> element_of_set(U::'a,intersection_of_members(Vf))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3491
\  (\\<forall>Vt X. topological_space(X::'a,Vt) --> equal_sets(union_of_members(Vt),X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3492
\  (\\<forall>X Vt. topological_space(X::'a,Vt) --> element_of_collection(empty_set::'a,Vt)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3493
\  (\\<forall>X Vt. topological_space(X::'a,Vt) --> element_of_collection(X::'a,Vt)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3494
\  (\\<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)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3495
\  (\\<forall>X Vf Vt. topological_space(X::'a,Vt) & subset_collections(Vf::'a,Vt) --> element_of_collection(union_of_members(Vf),Vt)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3496
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3497
\  (\\<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)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3498
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3499
\  (\\<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)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3500
\  (\\<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)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3501
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3502
\  (\\<forall>U X Vt. open(U::'a,X,Vt) --> topological_space(X::'a,Vt)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3503
\  (\\<forall>X U Vt. open(U::'a,X,Vt) --> element_of_collection(U::'a,Vt)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3504
\  (\\<forall>X U Vt. topological_space(X::'a,Vt) & element_of_collection(U::'a,Vt) --> open(U::'a,X,Vt)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3505
\  (\\<forall>U X Vt. closed(U::'a,X,Vt) --> topological_space(X::'a,Vt)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3506
\  (\\<forall>U X Vt. closed(U::'a,X,Vt) --> open(relative_complement_sets(U::'a,X),X,Vt)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3507
\  (\\<forall>U X Vt. topological_space(X::'a,Vt) & open(relative_complement_sets(U::'a,X),X,Vt) --> closed(U::'a,X,Vt)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3508
\  (\\<forall>Vs X Vt. finer(Vt::'a,Vs,X) --> topological_space(X::'a,Vt)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3509
\  (\\<forall>Vt X Vs. finer(Vt::'a,Vs,X) --> topological_space(X::'a,Vs)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3510
\  (\\<forall>X Vs Vt. finer(Vt::'a,Vs,X) --> subset_collections(Vs::'a,Vt)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3511
\  (\\<forall>X Vs Vt. topological_space(X::'a,Vt) & topological_space(X::'a,Vs) & subset_collections(Vs::'a,Vt) --> finer(Vt::'a,Vs,X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3512
\  (\\<forall>Vf X. basis(X::'a,Vf) --> equal_sets(union_of_members(Vf),X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3513
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3514
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3515
\  (\\<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))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3516
\  (\\<forall>Vf X. equal_sets(union_of_members(Vf),X) --> basis(X::'a,Vf) | element_of_set(f7(X::'a,Vf),X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3517
\  (\\<forall>X Vf. equal_sets(union_of_members(Vf),X) --> basis(X::'a,Vf) | element_of_collection(f8(X::'a,Vf),Vf)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3518
\  (\\<forall>X Vf. equal_sets(union_of_members(Vf),X) --> basis(X::'a,Vf) | element_of_collection(f9(X::'a,Vf),Vf)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3519
\  (\\<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)))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3520
\  (\\<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)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3521
\  (\\<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))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3522
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3523
\  (\\<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)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3524
\  (\\<forall>Vf U. element_of_collection(U::'a,top_of_basis(Vf)) | element_of_set(f11(Vf::'a,U),U)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3525
\  (\\<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))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3526
\  (\\<forall>U Y X Vt. element_of_collection(U::'a,subspace_topology(X::'a,Vt,Y)) --> topological_space(X::'a,Vt)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3527
\  (\\<forall>U Vt Y X. element_of_collection(U::'a,subspace_topology(X::'a,Vt,Y)) --> subset_sets(Y::'a,X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3528
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3529
\  (\\<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)))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3530
\  (\\<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))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3531
\  (\\<forall>U Y X Vt. element_of_set(U::'a,interior(Y::'a,X,Vt)) --> topological_space(X::'a,Vt)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3532
\  (\\<forall>U Vt Y X. element_of_set(U::'a,interior(Y::'a,X,Vt)) --> subset_sets(Y::'a,X)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3533
\  (\\<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))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3534
\  (\\<forall>X Vt U Y. element_of_set(U::'a,interior(Y::'a,X,Vt)) --> subset_sets(f13(Y::'a,X,Vt,U),Y)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3535
\  (\\<forall>Y U X Vt. element_of_set(U::'a,interior(Y::'a,X,Vt)) --> open(f13(Y::'a,X,Vt,U),X,Vt)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3536
\  (\\<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))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3537
\  (\\<forall>U Y X Vt. element_of_set(U::'a,closure(Y::'a,X,Vt)) --> topological_space(X::'a,Vt)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3538
\  (\\<forall>U Vt Y X. element_of_set(U::'a,closure(Y::'a,X,Vt)) --> subset_sets(Y::'a,X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3539
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3540
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3541
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3542
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3543
\  (\\<forall>U Y X Vt. neighborhood(U::'a,Y,X,Vt) --> topological_space(X::'a,Vt)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3544
\  (\\<forall>Y U X Vt. neighborhood(U::'a,Y,X,Vt) --> open(U::'a,X,Vt)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3545
\  (\\<forall>X Vt Y U. neighborhood(U::'a,Y,X,Vt) --> element_of_set(Y::'a,U)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3546
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3547
\  (\\<forall>Z Y X Vt. limit_point(Z::'a,Y,X,Vt) --> topological_space(X::'a,Vt)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3548
\  (\\<forall>Z Vt Y X. limit_point(Z::'a,Y,X,Vt) --> subset_sets(Y::'a,X)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3549
\  (\\<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))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3550
\  (\\<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))) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3551
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3552
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3553
\  (\\<forall>U Y X Vt. element_of_set(U::'a,boundary(Y::'a,X,Vt)) --> topological_space(X::'a,Vt)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3554
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3555
\  (\\<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))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3556
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3557
\  (\\<forall>X Vt. hausdorff(X::'a,Vt) --> topological_space(X::'a,Vt)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3558
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3559
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3560
\  (\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3561
\  (\\<forall>Vt X. topological_space(X::'a,Vt) --> hausdorff(X::'a,Vt) | element_of_set(f19(X::'a,Vt),X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3562
\  (\\<forall>Vt X. topological_space(X::'a,Vt) --> hausdorff(X::'a,Vt) | element_of_set(f20(X::'a,Vt),X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3563
\  (\\<forall>X Vt. topological_space(X::'a,Vt) & eq_p(f19(X::'a,Vt),f20(X::'a,Vt)) --> hausdorff(X::'a,Vt)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3564
\  (\\<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)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3565
\  (\\<forall>Va1 Va2 X Vt. separation(Va1::'a,Va2,X,Vt) --> topological_space(X::'a,Vt)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3566
\  (\\<forall>Va2 X Vt Va1. ~(separation(Va1::'a,Va2,X,Vt) & equal_sets(Va1::'a,empty_set))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3567
\  (\\<forall>Va1 X Vt Va2. ~(separation(Va1::'a,Va2,X,Vt) & equal_sets(Va2::'a,empty_set))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3568
\  (\\<forall>Va2 X Va1 Vt. separation(Va1::'a,Va2,X,Vt) --> element_of_collection(Va1::'a,Vt)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3569
\  (\\<forall>Va1 X Va2 Vt. separation(Va1::'a,Va2,X,Vt) --> element_of_collection(Va2::'a,Vt)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3570
\  (\\<forall>Vt Va1 Va2 X. separation(Va1::'a,Va2,X,Vt) --> equal_sets(union_of_sets(Va1::'a,Va2),X)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3571
\  (\\<forall>X Vt Va1 Va2. separation(Va1::'a,Va2,X,Vt) --> disjoint_s(Va1::'a,Va2)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3572
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3573
\  (\\<forall>X Vt. connected_space(X::'a,Vt) --> topological_space(X::'a,Vt)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3574
\  (\\<forall>Va1 Va2 X Vt. ~(connected_space(X::'a,Vt) & separation(Va1::'a,Va2,X,Vt))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3575
\  (\\<forall>X Vt. topological_space(X::'a,Vt) --> connected_space(X::'a,Vt) | separation(f21(X::'a,Vt),f22(X::'a,Vt),X,Vt)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3576
\  (\\<forall>Va X Vt. connected_set(Va::'a,X,Vt) --> topological_space(X::'a,Vt)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3577
\  (\\<forall>Vt Va X. connected_set(Va::'a,X,Vt) --> subset_sets(Va::'a,X)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3578
\  (\\<forall>X Vt Va. connected_set(Va::'a,X,Vt) --> connected_space(Va::'a,subspace_topology(X::'a,Vt,Va))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3579
\  (\\<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)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3580
\  (\\<forall>Vf X Vt. open_covering(Vf::'a,X,Vt) --> topological_space(X::'a,Vt)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3581
\  (\\<forall>X Vf Vt. open_covering(Vf::'a,X,Vt) --> subset_collections(Vf::'a,Vt)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3582
\  (\\<forall>Vt Vf X. open_covering(Vf::'a,X,Vt) --> equal_sets(union_of_members(Vf),X)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3583
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3584
\  (\\<forall>X Vt. compact_space(X::'a,Vt) --> topological_space(X::'a,Vt)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3585
\  (\\<forall>X Vt Vf1. compact_space(X::'a,Vt) & open_covering(Vf1::'a,X,Vt) --> finite(f23(X::'a,Vt,Vf1))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3586
\  (\\<forall>X Vt Vf1. compact_space(X::'a,Vt) & open_covering(Vf1::'a,X,Vt) --> subset_collections(f23(X::'a,Vt,Vf1),Vf1)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3587
\  (\\<forall>Vf1 X Vt. compact_space(X::'a,Vt) & open_covering(Vf1::'a,X,Vt) --> open_covering(f23(X::'a,Vt,Vf1),X,Vt)) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3588
\  (\\<forall>X Vt. topological_space(X::'a,Vt) --> compact_space(X::'a,Vt) | open_covering(f24(X::'a,Vt),X,Vt)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3589
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3590
\  (\\<forall>Va X Vt. compact_set(Va::'a,X,Vt) --> topological_space(X::'a,Vt)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3591
\  (\\<forall>Vt Va X. compact_set(Va::'a,X,Vt) --> subset_sets(Va::'a,X)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3592
\  (\\<forall>X Vt Va. compact_set(Va::'a,X,Vt) --> compact_space(Va::'a,subspace_topology(X::'a,Vt,Va))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3593
\  (\\<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)) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3594
\  (basis(cx::'a,f)) &      \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3595
\  (\\<forall>U. element_of_collection(U::'a,top_of_basis(f))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3596
\  (\\<forall>V. element_of_collection(V::'a,top_of_basis(f))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3597
\  (\\<forall>U V. ~element_of_collection(intersection_of_sets(U::'a,V),top_of_basis(f))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3598
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3599
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3600
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3601
(*0 inferences so far.  Searching to depth 0.  0.8 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3602
val TOP004_2 = prove
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3603
 ("(\\<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))) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3604
\  (\\<forall>Vf X. basis(X::'a,Vf) --> equal_sets(union_of_members(Vf),X)) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3605
\  (\\<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))) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3606
\  (\\<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)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3607
\  (\\<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))) &  \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3608
\  (\\<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))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3609
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3610
\  (\\<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)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3611
\  (\\<forall>Vf U. element_of_collection(U::'a,top_of_basis(Vf)) | element_of_set(f11(Vf::'a,U),U)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3612
\  (\\<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))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3613
\  (\\<forall>Y X Z. subset_sets(X::'a,Y) & subset_sets(Y::'a,Z) --> subset_sets(X::'a,Z)) &        \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3614
\  (\\<forall>Y Z X. element_of_set(Z::'a,intersection_of_sets(X::'a,Y)) --> element_of_set(Z::'a,X)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3615
\  (\\<forall>X Z Y. element_of_set(Z::'a,intersection_of_sets(X::'a,Y)) --> element_of_set(Z::'a,Y)) &     \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3616
\  (\\<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))) &       \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3617
\  (\\<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))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3618
\  (\\<forall>X Z Y. equal_sets(X::'a,Y) & element_of_set(Z::'a,X) --> element_of_set(Z::'a,Y)) &   \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3619
\  (\\<forall>Y X. equal_sets(intersection_of_sets(X::'a,Y),intersection_of_sets(Y::'a,X))) &   \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3620
\  (basis(cx::'a,f)) &      \
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3621
\  (\\<forall>U. element_of_collection(U::'a,top_of_basis(f))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3622
\  (\\<forall>V. element_of_collection(V::'a,top_of_basis(f))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3623
\  (\\<forall>U V. ~element_of_collection(intersection_of_sets(U::'a,V),top_of_basis(f))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3624
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3625
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3626
(*53777 inferences so far.  Searching to depth 20.  68.7 secs*)
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3627
val TOP005_2 = prove_hard
14183
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3628
 ("(\\<forall>Vf U. element_of_set(U::'a,union_of_members(Vf)) --> element_of_set(U::'a,f1(Vf::'a,U))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3629
\  (\\<forall>U Vf. element_of_set(U::'a,union_of_members(Vf)) --> element_of_collection(f1(Vf::'a,U),Vf)) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3630
\  (\\<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))) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3631
\  (\\<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)) &      \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3632
\  (\\<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)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3633
\  (\\<forall>Vf U. element_of_collection(U::'a,top_of_basis(Vf)) | element_of_set(f11(Vf::'a,U),U)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3634
\  (\\<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))) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3635
\  (\\<forall>X U Y. element_of_set(U::'a,X) --> subset_sets(X::'a,Y) | element_of_set(U::'a,Y)) & \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3636
\  (\\<forall>Y X Z. subset_sets(X::'a,Y) & element_of_collection(Y::'a,Z) --> subset_sets(X::'a,union_of_members(Z))) &    \
466a2a69e7e8 quantifier symbols
paulson
parents: 11451
diff changeset
  3637
\  (\\<forall>X U Y. subset_collections(X::'a,Y) & element_of_collection(U::'a,X) --> element_of_collection(U::'a,Y)) &     \
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3638
\  (subset_collections(g::'a,top_of_basis(f))) &    \
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3639
\  (~element_of_collection(union_of_members(g),top_of_basis(f))) --> False",
9841
ca3173f87b5c safe_meson_tac -> meson_tac
paulson
parents: 9000
diff changeset
  3640
  meson_tac 1);
8557
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3641
fe75fe482566 restored the MESON examples file HOL/ex/mesontest2.ML
paulson
parents:
diff changeset
  3642