| author | wenzelm | 
| Thu, 24 Mar 2011 16:47:24 +0100 | |
| changeset 42082 | 47f8bfe0f597 | 
| parent 41144 | 509e51b7509a | 
| child 42103 | 6066a35f6678 | 
| permissions | -rw-r--r-- | 
| 33027 | 1  | 
(* Title: HOL/Metis_Examples/set.thy  | 
| 23449 | 2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 41144 | 3  | 
Author: Jasmin Blanchette, TU Muenchen  | 
| 23449 | 4  | 
|
| 41144 | 5  | 
Testing Metis.  | 
| 23449 | 6  | 
*)  | 
7  | 
||
| 33027 | 8  | 
theory set  | 
9  | 
imports Main  | 
|
| 23449 | 10  | 
begin  | 
11  | 
||
12  | 
lemma "EX x X. ALL y. EX z Z. (~P(y,y) | P(x,x) | ~S(z,x)) &  | 
|
13  | 
(S(x,y) | ~S(y,z) | Q(Z,Z)) &  | 
|
| 
24742
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
23519 
diff
changeset
 | 
14  | 
(Q(X,y) | ~Q(y,Z) | S(X,X))"  | 
| 23519 | 15  | 
by metis  | 
| 23449 | 16  | 
|
17  | 
lemma "P(n::nat) ==> ~P(0) ==> n ~= 0"  | 
|
18  | 
by metis  | 
|
19  | 
||
| 36925 | 20  | 
sledgehammer_params [isar_proof, isar_shrink_factor = 1]  | 
| 23449 | 21  | 
|
22  | 
(*multiple versions of this example*)  | 
|
23  | 
lemma (*equal_union: *)  | 
|
| 36566 | 24  | 
"(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"  | 
| 
37325
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
25  | 
proof -  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
26  | 
have F1: "\<forall>(x\<^isub>2\<Colon>'b \<Rightarrow> bool) x\<^isub>1\<Colon>'b \<Rightarrow> bool. x\<^isub>1 \<subseteq> x\<^isub>1 \<union> x\<^isub>2" by (metis Un_commute Un_upper2)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
27  | 
have F2a: "\<forall>(x\<^isub>2\<Colon>'b \<Rightarrow> bool) x\<^isub>1\<Colon>'b \<Rightarrow> bool. x\<^isub>1 \<subseteq> x\<^isub>2 \<longrightarrow> x\<^isub>2 = x\<^isub>2 \<union> x\<^isub>1" by (metis Un_commute subset_Un_eq)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
28  | 
have F2: "\<forall>(x\<^isub>2\<Colon>'b \<Rightarrow> bool) x\<^isub>1\<Colon>'b \<Rightarrow> bool. x\<^isub>1 \<subseteq> x\<^isub>2 \<and> x\<^isub>2 \<subseteq> x\<^isub>1 \<longrightarrow> x\<^isub>1 = x\<^isub>2" by (metis F2a subset_Un_eq)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
29  | 
  { assume "\<not> Z \<subseteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
30  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
31  | 
moreover  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
32  | 
  { assume AA1: "Y \<union> Z \<noteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
33  | 
    { assume "\<not> Y \<subseteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
34  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
35  | 
moreover  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
36  | 
    { assume AAA1: "Y \<subseteq> X \<and> Y \<union> Z \<noteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
37  | 
      { assume "\<not> Z \<subseteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
38  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
39  | 
moreover  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
40  | 
      { assume "(Z \<subseteq> X \<and> Y \<subseteq> X) \<and> Y \<union> Z \<noteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
41  | 
hence "Y \<union> Z \<subseteq> X \<and> X \<noteq> Y \<union> Z" by (metis Un_subset_iff)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
42  | 
hence "Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> Y \<union> Z" by (metis F2)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
43  | 
hence "\<exists>x\<^isub>1\<Colon>'a \<Rightarrow> bool. Y \<subseteq> x\<^isub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^isub>1 \<union> Z" by (metis F1)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
44  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
45  | 
ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AAA1) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
46  | 
ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
47  | 
moreover  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
48  | 
  { assume "\<exists>x\<^isub>1\<Colon>'a \<Rightarrow> bool. (Z \<subseteq> x\<^isub>1 \<and> Y \<subseteq> x\<^isub>1) \<and> \<not> X \<subseteq> x\<^isub>1"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
49  | 
    { assume "\<not> Y \<subseteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
50  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
51  | 
moreover  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
52  | 
    { assume AAA1: "Y \<subseteq> X \<and> Y \<union> Z \<noteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
53  | 
      { assume "\<not> Z \<subseteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
54  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
55  | 
moreover  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
56  | 
      { assume "(Z \<subseteq> X \<and> Y \<subseteq> X) \<and> Y \<union> Z \<noteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
57  | 
hence "Y \<union> Z \<subseteq> X \<and> X \<noteq> Y \<union> Z" by (metis Un_subset_iff)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
58  | 
hence "Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> Y \<union> Z" by (metis F2)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
59  | 
hence "\<exists>x\<^isub>1\<Colon>'a \<Rightarrow> bool. Y \<subseteq> x\<^isub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^isub>1 \<union> Z" by (metis F1)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
60  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
61  | 
ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AAA1) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
62  | 
ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by blast }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
63  | 
moreover  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
64  | 
  { assume "\<not> Y \<subseteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
65  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
66  | 
ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by metis  | 
| 23449 | 67  | 
qed  | 
68  | 
||
| 36925 | 69  | 
sledgehammer_params [isar_proof, isar_shrink_factor = 2]  | 
| 23449 | 70  | 
|
71  | 
lemma (*equal_union: *)  | 
|
| 
37325
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
72  | 
"(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
73  | 
proof -  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
74  | 
have F1: "\<forall>(x\<^isub>2\<Colon>'b \<Rightarrow> bool) x\<^isub>1\<Colon>'b \<Rightarrow> bool. x\<^isub>1 \<subseteq> x\<^isub>2 \<and> x\<^isub>2 \<subseteq> x\<^isub>1 \<longrightarrow> x\<^isub>1 = x\<^isub>2" by (metis Un_commute subset_Un_eq)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
75  | 
  { assume AA1: "\<exists>x\<^isub>1\<Colon>'a \<Rightarrow> bool. (Z \<subseteq> x\<^isub>1 \<and> Y \<subseteq> x\<^isub>1) \<and> \<not> X \<subseteq> x\<^isub>1"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
76  | 
    { assume AAA1: "Y \<subseteq> X \<and> Y \<union> Z \<noteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
77  | 
      { assume "\<not> Z \<subseteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
78  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
79  | 
moreover  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
80  | 
      { assume "Y \<union> Z \<subseteq> X \<and> X \<noteq> Y \<union> Z"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
81  | 
hence "\<exists>x\<^isub>1\<Colon>'a \<Rightarrow> bool. Y \<subseteq> x\<^isub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^isub>1 \<union> Z" by (metis F1 Un_commute Un_upper2)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
82  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
83  | 
ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AAA1 Un_subset_iff) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
84  | 
moreover  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
85  | 
    { assume "\<not> Y \<subseteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
86  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
87  | 
ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 Un_subset_iff) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
88  | 
moreover  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
89  | 
  { assume "\<not> Z \<subseteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
90  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
91  | 
moreover  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
92  | 
  { assume "\<not> Y \<subseteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
93  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
94  | 
moreover  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
95  | 
  { assume AA1: "Y \<subseteq> X \<and> Y \<union> Z \<noteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
96  | 
    { assume "\<not> Z \<subseteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
97  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
98  | 
moreover  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
99  | 
    { assume "Y \<union> Z \<subseteq> X \<and> X \<noteq> Y \<union> Z"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
100  | 
hence "\<exists>x\<^isub>1\<Colon>'a \<Rightarrow> bool. Y \<subseteq> x\<^isub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^isub>1 \<union> Z" by (metis F1 Un_commute Un_upper2)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
101  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
102  | 
ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 Un_subset_iff) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
103  | 
ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by metis  | 
| 23449 | 104  | 
qed  | 
105  | 
||
| 36925 | 106  | 
sledgehammer_params [isar_proof, isar_shrink_factor = 3]  | 
| 23449 | 107  | 
|
108  | 
lemma (*equal_union: *)  | 
|
| 
37325
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
109  | 
"(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
110  | 
proof -  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
111  | 
have F1a: "\<forall>(x\<^isub>2\<Colon>'b \<Rightarrow> bool) x\<^isub>1\<Colon>'b \<Rightarrow> bool. x\<^isub>1 \<subseteq> x\<^isub>2 \<longrightarrow> x\<^isub>2 = x\<^isub>2 \<union> x\<^isub>1" by (metis Un_commute subset_Un_eq)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
112  | 
have F1: "\<forall>(x\<^isub>2\<Colon>'b \<Rightarrow> bool) x\<^isub>1\<Colon>'b \<Rightarrow> bool. x\<^isub>1 \<subseteq> x\<^isub>2 \<and> x\<^isub>2 \<subseteq> x\<^isub>1 \<longrightarrow> x\<^isub>1 = x\<^isub>2" by (metis F1a subset_Un_eq)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
113  | 
  { assume "(Z \<subseteq> X \<and> Y \<subseteq> X) \<and> Y \<union> Z \<noteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
114  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1 Un_commute Un_subset_iff Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
115  | 
moreover  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
116  | 
  { assume AA1: "\<exists>x\<^isub>1\<Colon>'a \<Rightarrow> bool. (Z \<subseteq> x\<^isub>1 \<and> Y \<subseteq> x\<^isub>1) \<and> \<not> X \<subseteq> x\<^isub>1"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
117  | 
    { assume "(Z \<subseteq> X \<and> Y \<subseteq> X) \<and> Y \<union> Z \<noteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
118  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1 Un_commute Un_subset_iff Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
119  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 Un_commute Un_subset_iff Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
120  | 
ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2)  | 
| 23449 | 121  | 
qed  | 
122  | 
||
| 36925 | 123  | 
sledgehammer_params [isar_proof, isar_shrink_factor = 4]  | 
| 23449 | 124  | 
|
125  | 
lemma (*equal_union: *)  | 
|
| 
37325
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
126  | 
"(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
127  | 
proof -  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
128  | 
have F1: "\<forall>(x\<^isub>2\<Colon>'b \<Rightarrow> bool) x\<^isub>1\<Colon>'b \<Rightarrow> bool. x\<^isub>1 \<subseteq> x\<^isub>2 \<and> x\<^isub>2 \<subseteq> x\<^isub>1 \<longrightarrow> x\<^isub>1 = x\<^isub>2" by (metis Un_commute subset_Un_eq)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
129  | 
  { assume "\<not> Y \<subseteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
130  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
131  | 
moreover  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
132  | 
  { assume AA1: "Y \<subseteq> X \<and> Y \<union> Z \<noteq> X"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
133  | 
    { assume "\<exists>x\<^isub>1\<Colon>'a \<Rightarrow> bool. Y \<subseteq> x\<^isub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^isub>1 \<union> Z"
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
134  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
135  | 
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 F1 Un_commute Un_subset_iff Un_upper2) }  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
136  | 
ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a \<Rightarrow> bool. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_subset_iff Un_upper2)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
137  | 
qed  | 
| 23449 | 138  | 
|
| 
37325
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
139  | 
sledgehammer_params [isar_proof, isar_shrink_factor = 1]  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
140  | 
|
| 23449 | 141  | 
lemma (*equal_union: *)  | 
| 
37325
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
142  | 
"(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"  | 
| 23449 | 143  | 
by (metis Un_least Un_upper1 Un_upper2 set_eq_subset)  | 
144  | 
||
| 
37325
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
145  | 
lemma "(X = Y \<inter> Z) = (X \<subseteq> Y \<and> X \<subseteq> Z \<and> (\<forall>V. V \<subseteq> Y \<and> V \<subseteq> Z \<longrightarrow> V \<subseteq> X))"  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
146  | 
by (metis Int_greatest Int_lower1 Int_lower2 subset_antisym)  | 
| 23449 | 147  | 
|
| 
37325
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
148  | 
lemma fixedpoint: "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y"  | 
| 23449 | 149  | 
by metis  | 
150  | 
||
| 
37325
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
151  | 
lemma (* fixedpoint: *) "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y"  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
152  | 
proof -  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
153  | 
assume "\<exists>!x\<Colon>'a. f (g x) = x"  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
154  | 
thus "\<exists>!y\<Colon>'b. g (f y) = y" by metis  | 
| 23449 | 155  | 
qed  | 
156  | 
||
| 
37325
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
157  | 
lemma (* singleton_example_2: *)  | 
| 23449 | 158  | 
     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
 | 
159  | 
by (metis Set.subsetI Union_upper insertCI set_eq_subset)  | 
|
160  | 
||
| 
37325
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
161  | 
lemma (* singleton_example_2: *)  | 
| 23449 | 162  | 
     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
 | 
| 32685 | 163  | 
by (metis Set.subsetI Union_upper insert_iff set_eq_subset)  | 
| 23449 | 164  | 
|
165  | 
lemma singleton_example_2:  | 
|
166  | 
     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
 | 
|
| 
37325
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
167  | 
proof -  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
168  | 
assume "\<forall>x \<in> S. \<Union>S \<subseteq> x"  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
169  | 
hence "\<forall>x\<^isub>1. x\<^isub>1 \<subseteq> \<Union>S \<and> x\<^isub>1 \<in> S \<longrightarrow> x\<^isub>1 = \<Union>S" by (metis set_eq_subset)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
170  | 
hence "\<forall>x\<^isub>1. x\<^isub>1 \<in> S \<longrightarrow> x\<^isub>1 = \<Union>S" by (metis Union_upper)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
171  | 
  hence "\<forall>x\<^isub>1\<Colon>('a \<Rightarrow> bool) \<Rightarrow> bool. \<Union>S \<in> x\<^isub>1 \<longrightarrow> S \<subseteq> x\<^isub>1" by (metis subsetI)
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
172  | 
  hence "\<forall>x\<^isub>1\<Colon>('a \<Rightarrow> bool) \<Rightarrow> bool. S \<subseteq> insert (\<Union>S) x\<^isub>1" by (metis insert_iff)
 | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
173  | 
  thus "\<exists>z. S \<subseteq> {z}" by metis
 | 
| 23449 | 174  | 
qed  | 
175  | 
||
176  | 
text {*
 | 
|
177  | 
From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages  | 
|
178  | 
293-314.  | 
|
179  | 
*}  | 
|
180  | 
||
| 
37325
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
181  | 
(* Notes: (1) The numbering doesn't completely agree with the paper.  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
182  | 
(2) We must rename set variables to avoid type clashes. *)  | 
| 23449 | 183  | 
lemma "\<exists>B. (\<forall>x \<in> B. x \<le> (0::int))"  | 
184  | 
"D \<in> F \<Longrightarrow> \<exists>G. \<forall>A \<in> G. \<exists>B \<in> F. A \<subseteq> B"  | 
|
185  | 
"P a \<Longrightarrow> \<exists>A. (\<forall>x \<in> A. P x) \<and> (\<exists>y. y \<in> A)"  | 
|
186  | 
"a < b \<and> b < (c::int) \<Longrightarrow> \<exists>B. a \<notin> B \<and> b \<in> B \<and> c \<notin> B"  | 
|
187  | 
"P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"  | 
|
188  | 
"P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"  | 
|
189  | 
"\<exists>A. a \<notin> A"  | 
|
| 36566 | 190  | 
"(\<forall>C. (0, 0) \<in> C \<and> (\<forall>x y. (x, y) \<in> C \<longrightarrow> (Suc x, Suc y) \<in> C) \<longrightarrow> (n, m) \<in> C) \<and> Q n \<longrightarrow> Q m"  | 
| 
37325
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
191  | 
apply (metis all_not_in_conv)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
192  | 
apply (metis all_not_in_conv)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
193  | 
apply (metis mem_def)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
194  | 
apply (metis less_int_def singleton_iff)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
195  | 
apply (metis mem_def)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
196  | 
apply (metis mem_def)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
197  | 
apply (metis all_not_in_conv)  | 
| 
 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 
blanchet 
parents: 
36925 
diff
changeset
 | 
198  | 
by (metis pair_in_Id_conv)  | 
| 23449 | 199  | 
|
200  | 
end  |