| author | wenzelm | 
| Mon, 06 Apr 2015 23:14:05 +0200 | |
| changeset 59940 | 087d81f5213e | 
| parent 58889 | 5b7a9633cfa8 | 
| child 61076 | bdc1e2f0a86a | 
| permissions | -rw-r--r-- | 
| 43197 | 1 | (* Title: HOL/Metis_Examples/Sets.thy | 
| 2 | Author: Lawrence C. Paulson, Cambridge University Computer Laboratory | |
| 41144 | 3 | Author: Jasmin Blanchette, TU Muenchen | 
| 23449 | 4 | |
| 43197 | 5 | Metis example featuring typed set theory. | 
| 23449 | 6 | *) | 
| 7 | ||
| 58889 | 8 | section {* Metis Example Featuring Typed Set Theory *}
 | 
| 43197 | 9 | |
| 10 | theory Sets | |
| 33027 | 11 | imports Main | 
| 23449 | 12 | begin | 
| 13 | ||
| 50705 
0e943b33d907
use new skolemizer for reconstructing skolemization steps in Isar proofs (because the old skolemizer messes up the order of the Skolem arguments)
 blanchet parents: 
50020diff
changeset | 14 | declare [[metis_new_skolem]] | 
| 42103 
6066a35f6678
Metis examples use the new Skolemizer to test it
 blanchet parents: 
41144diff
changeset | 15 | |
| 23449 | 16 | lemma "EX x X. ALL y. EX z Z. (~P(y,y) | P(x,x) | ~S(z,x)) & | 
| 17 | (S(x,y) | ~S(y,z) | Q(Z,Z)) & | |
| 43197 | 18 | (Q(X,y) | ~Q(y,Z) | S(X,X))" | 
| 23519 | 19 | by metis | 
| 23449 | 20 | |
| 21 | lemma "P(n::nat) ==> ~P(0) ==> n ~= 0" | |
| 22 | by metis | |
| 23 | ||
| 57245 | 24 | sledgehammer_params [isar_proofs, compress = 1] | 
| 23449 | 25 | |
| 26 | (*multiple versions of this example*) | |
| 46077 | 27 | lemma (*equal_union: *) | 
| 36566 | 28 | "(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: 
36925diff
changeset | 29 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 30 | have F1: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>1 \<union> x\<^sub>2" by (metis Un_commute Un_upper2) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 31 | have F2a: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<longrightarrow> x\<^sub>2 = x\<^sub>2 \<union> x\<^sub>1" by (metis Un_commute subset_Un_eq) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 32 | have F2: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<and> x\<^sub>2 \<subseteq> x\<^sub>1 \<longrightarrow> x\<^sub>1 = x\<^sub>2" by (metis F2a subset_Un_eq) | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 33 |   { assume "\<not> Z \<subseteq> X"
 | 
| 46077 | 34 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) } | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 35 | moreover | 
| 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 36 |   { 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: 
36925diff
changeset | 37 |     { assume "\<not> Y \<subseteq> X"
 | 
| 46077 | 38 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1) } | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 39 | moreover | 
| 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 40 |     { 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: 
36925diff
changeset | 41 |       { assume "\<not> Z \<subseteq> X"
 | 
| 46077 | 42 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) } | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 43 | moreover | 
| 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 44 |       { 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: 
36925diff
changeset | 45 | 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: 
36925diff
changeset | 46 | hence "Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> Y \<union> Z" by (metis F2) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 47 | hence "\<exists>x\<^sub>1\<Colon>'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z" by (metis F1) | 
| 46077 | 48 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) } | 
| 49 | ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AAA1) } | |
| 50 | ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1) } | |
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 51 | moreover | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 52 |   { assume "\<exists>x\<^sub>1\<Colon>'a set. (Z \<subseteq> x\<^sub>1 \<and> Y \<subseteq> x\<^sub>1) \<and> \<not> X \<subseteq> x\<^sub>1"
 | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 53 |     { assume "\<not> Y \<subseteq> X"
 | 
| 46077 | 54 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1) } | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 55 | moreover | 
| 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 56 |     { 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: 
36925diff
changeset | 57 |       { assume "\<not> Z \<subseteq> X"
 | 
| 46077 | 58 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) } | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 59 | moreover | 
| 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 60 |       { 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: 
36925diff
changeset | 61 | 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: 
36925diff
changeset | 62 | hence "Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> Y \<union> Z" by (metis F2) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 63 | hence "\<exists>x\<^sub>1\<Colon>'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z" by (metis F1) | 
| 46077 | 64 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) } | 
| 65 | ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AAA1) } | |
| 66 | ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by blast } | |
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 67 | moreover | 
| 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 68 |   { assume "\<not> Y \<subseteq> X"
 | 
| 46077 | 69 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1) } | 
| 70 | ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by metis | |
| 23449 | 71 | qed | 
| 72 | ||
| 57245 | 73 | sledgehammer_params [isar_proofs, compress = 2] | 
| 23449 | 74 | |
| 75 | lemma (*equal_union: *) | |
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 76 | "(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: 
36925diff
changeset | 77 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 78 | have F1: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<and> x\<^sub>2 \<subseteq> x\<^sub>1 \<longrightarrow> x\<^sub>1 = x\<^sub>2" by (metis Un_commute subset_Un_eq) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 79 |   { assume AA1: "\<exists>x\<^sub>1\<Colon>'a set. (Z \<subseteq> x\<^sub>1 \<and> Y \<subseteq> x\<^sub>1) \<and> \<not> X \<subseteq> x\<^sub>1"
 | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 80 |     { 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: 
36925diff
changeset | 81 |       { assume "\<not> Z \<subseteq> X"
 | 
| 46077 | 82 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) } | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 83 | moreover | 
| 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 84 |       { assume "Y \<union> Z \<subseteq> X \<and> X \<noteq> Y \<union> Z"
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 85 | hence "\<exists>x\<^sub>1\<Colon>'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z" by (metis F1 Un_commute Un_upper2) | 
| 46077 | 86 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) } | 
| 87 | ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AAA1 Un_subset_iff) } | |
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 88 | moreover | 
| 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 89 |     { assume "\<not> Y \<subseteq> X"
 | 
| 46077 | 90 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2) } | 
| 91 | ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 Un_subset_iff) } | |
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 92 | moreover | 
| 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 93 |   { assume "\<not> Z \<subseteq> X"
 | 
| 46077 | 94 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) } | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 95 | moreover | 
| 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 96 |   { assume "\<not> Y \<subseteq> X"
 | 
| 46077 | 97 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2) } | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 98 | moreover | 
| 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 99 |   { 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: 
36925diff
changeset | 100 |     { assume "\<not> Z \<subseteq> X"
 | 
| 46077 | 101 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) } | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 102 | moreover | 
| 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 103 |     { assume "Y \<union> Z \<subseteq> X \<and> X \<noteq> Y \<union> Z"
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 104 | hence "\<exists>x\<^sub>1\<Colon>'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z" by (metis F1 Un_commute Un_upper2) | 
| 46077 | 105 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) } | 
| 106 | ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 Un_subset_iff) } | |
| 107 | ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by metis | |
| 23449 | 108 | qed | 
| 109 | ||
| 57245 | 110 | sledgehammer_params [isar_proofs, compress = 3] | 
| 23449 | 111 | |
| 112 | lemma (*equal_union: *) | |
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 113 | "(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: 
36925diff
changeset | 114 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 115 | have F1a: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<longrightarrow> x\<^sub>2 = x\<^sub>2 \<union> x\<^sub>1" by (metis Un_commute subset_Un_eq) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 116 | have F1: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<and> x\<^sub>2 \<subseteq> x\<^sub>1 \<longrightarrow> x\<^sub>1 = x\<^sub>2" by (metis F1a subset_Un_eq) | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 117 |   { assume "(Z \<subseteq> X \<and> Y \<subseteq> X) \<and> Y \<union> Z \<noteq> X"
 | 
| 46077 | 118 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1 Un_commute Un_subset_iff Un_upper2) } | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 119 | moreover | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 120 |   { assume AA1: "\<exists>x\<^sub>1\<Colon>'a set. (Z \<subseteq> x\<^sub>1 \<and> Y \<subseteq> x\<^sub>1) \<and> \<not> X \<subseteq> x\<^sub>1"
 | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 121 |     { assume "(Z \<subseteq> X \<and> Y \<subseteq> X) \<and> Y \<union> Z \<noteq> X"
 | 
| 46077 | 122 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1 Un_commute Un_subset_iff Un_upper2) } | 
| 123 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 Un_commute Un_subset_iff Un_upper2) } | |
| 124 | ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2) | |
| 23449 | 125 | qed | 
| 126 | ||
| 57245 | 127 | sledgehammer_params [isar_proofs, compress = 4] | 
| 23449 | 128 | |
| 129 | lemma (*equal_union: *) | |
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 130 | "(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: 
36925diff
changeset | 131 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 132 | have F1: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<and> x\<^sub>2 \<subseteq> x\<^sub>1 \<longrightarrow> x\<^sub>1 = x\<^sub>2" by (metis Un_commute subset_Un_eq) | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 133 |   { assume "\<not> Y \<subseteq> X"
 | 
| 46077 | 134 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2) } | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 135 | moreover | 
| 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 136 |   { assume AA1: "Y \<subseteq> X \<and> Y \<union> Z \<noteq> X"
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 137 |     { assume "\<exists>x\<^sub>1\<Colon>'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z"
 | 
| 46077 | 138 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) } | 
| 139 | hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 F1 Un_commute Un_subset_iff Un_upper2) } | |
| 140 | ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_subset_iff Un_upper2) | |
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 141 | qed | 
| 23449 | 142 | |
| 57245 | 143 | sledgehammer_params [isar_proofs, compress = 1] | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 144 | |
| 23449 | 145 | lemma (*equal_union: *) | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 146 | "(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 | 147 | by (metis Un_least Un_upper1 Un_upper2 set_eq_subset) | 
| 148 | ||
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 149 | 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: 
36925diff
changeset | 150 | by (metis Int_greatest Int_lower1 Int_lower2 subset_antisym) | 
| 23449 | 151 | |
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 152 | lemma fixedpoint: "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y" | 
| 23449 | 153 | by metis | 
| 154 | ||
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 155 | 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: 
36925diff
changeset | 156 | proof - | 
| 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 157 | 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: 
36925diff
changeset | 158 | thus "\<exists>!y\<Colon>'b. g (f y) = y" by metis | 
| 23449 | 159 | qed | 
| 160 | ||
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 161 | lemma (* singleton_example_2: *) | 
| 23449 | 162 |      "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
 | 
| 163 | by (metis Set.subsetI Union_upper insertCI set_eq_subset) | |
| 164 | ||
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 165 | lemma (* singleton_example_2: *) | 
| 23449 | 166 |      "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
 | 
| 32685 | 167 | by (metis Set.subsetI Union_upper insert_iff set_eq_subset) | 
| 23449 | 168 | |
| 169 | lemma singleton_example_2: | |
| 170 |      "\<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: 
36925diff
changeset | 171 | proof - | 
| 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 172 | assume "\<forall>x \<in> S. \<Union>S \<subseteq> x" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 173 | hence "\<forall>x\<^sub>1. x\<^sub>1 \<subseteq> \<Union>S \<and> x\<^sub>1 \<in> S \<longrightarrow> x\<^sub>1 = \<Union>S" by (metis set_eq_subset) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 174 | hence "\<forall>x\<^sub>1. x\<^sub>1 \<in> S \<longrightarrow> x\<^sub>1 = \<Union>S" by (metis Union_upper) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 175 |   hence "\<forall>x\<^sub>1\<Colon>('a set) set. \<Union>S \<in> x\<^sub>1 \<longrightarrow> S \<subseteq> x\<^sub>1" by (metis subsetI)
 | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51130diff
changeset | 176 |   hence "\<forall>x\<^sub>1\<Colon>('a set) set. S \<subseteq> insert (\<Union>S) x\<^sub>1" by (metis insert_iff)
 | 
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 177 |   thus "\<exists>z. S \<subseteq> {z}" by metis
 | 
| 23449 | 178 | qed | 
| 179 | ||
| 180 | text {*
 | |
| 181 | From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages | |
| 182 | 293-314. | |
| 183 | *} | |
| 184 | ||
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 185 | (* 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: 
36925diff
changeset | 186 | (2) We must rename set variables to avoid type clashes. *) | 
| 23449 | 187 | lemma "\<exists>B. (\<forall>x \<in> B. x \<le> (0::int))" | 
| 188 | "D \<in> F \<Longrightarrow> \<exists>G. \<forall>A \<in> G. \<exists>B \<in> F. A \<subseteq> B" | |
| 189 | "P a \<Longrightarrow> \<exists>A. (\<forall>x \<in> A. P x) \<and> (\<exists>y. y \<in> A)" | |
| 190 | "a < b \<and> b < (c::int) \<Longrightarrow> \<exists>B. a \<notin> B \<and> b \<in> B \<and> c \<notin> B" | |
| 191 | "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A" | |
| 192 | "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A" | |
| 193 | "\<exists>A. a \<notin> A" | |
| 36566 | 194 | "(\<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: 
36925diff
changeset | 195 | 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: 
36925diff
changeset | 196 | apply (metis all_not_in_conv) | 
| 46077 | 197 | apply (metis mem_Collect_eq) | 
| 48050 
72acba14c12b
definition less_int_def has changed, use 'less_le' instead
 huffman parents: 
46077diff
changeset | 198 | apply (metis less_le singleton_iff) | 
| 46077 | 199 | apply (metis mem_Collect_eq) | 
| 200 | apply (metis mem_Collect_eq) | |
| 37325 
c2a44bc874f9
redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
 blanchet parents: 
36925diff
changeset | 201 | 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: 
36925diff
changeset | 202 | by (metis pair_in_Id_conv) | 
| 23449 | 203 | |
| 204 | end |