src/HOL/ex/Set_Theory.thy
author wenzelm
Mon, 29 May 2017 19:34:07 +0200
changeset 65959 47309113ee4d
parent 63804 70554522bf98
child 67443 3abf6a722518
permissions -rw-r--r--
clarified view column; tuned whitespace;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
44276
fe769a0fcc96 avoid case-sensitive name for example theory
haftmann
parents: 41460
diff changeset
     1
(*  Title:      HOL/ex/Set_Theory.thy
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
     2
    Author:     Tobias Nipkow and Lawrence C Paulson
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
     3
    Copyright   1991  University of Cambridge
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
     4
*)
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
     5
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
     6
section \<open>Set Theory examples: Cantor's Theorem, Schröder-Bernstein Theorem, etc.\<close>
9100
9e081c812338 fixed deps;
wenzelm
parents:
diff changeset
     7
44276
fe769a0fcc96 avoid case-sensitive name for example theory
haftmann
parents: 41460
diff changeset
     8
theory Set_Theory
fe769a0fcc96 avoid case-sensitive name for example theory
haftmann
parents: 41460
diff changeset
     9
imports Main
fe769a0fcc96 avoid case-sensitive name for example theory
haftmann
parents: 41460
diff changeset
    10
begin
9100
9e081c812338 fixed deps;
wenzelm
parents:
diff changeset
    11
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
    12
text\<open>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    13
  These two are cited in Benzmueller and Kohlhase's system description
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    14
  of LEO, CADE-15, 1998 (pages 139-143) as theorems LEO could not
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    15
  prove.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
    16
\<close>
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    17
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    18
lemma "(X = Y \<union> Z) =
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    19
    (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    20
  by blast
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    21
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    22
lemma "(X = Y \<inter> Z) =
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    23
    (X \<subseteq> Y \<and> X \<subseteq> Z \<and> (\<forall>V. V \<subseteq> Y \<and> V \<subseteq> Z \<longrightarrow> V \<subseteq> X))"
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    24
  by blast
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    25
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
    26
text \<open>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    27
  Trivial example of term synthesis: apparently hard for some provers!
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
    28
\<close>
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    29
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
    30
schematic_goal "a \<noteq> b \<Longrightarrow> a \<in> ?X \<and> b \<notin> ?X"
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    31
  by blast
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    32
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    33
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    34
subsection \<open>Examples for the \<open>blast\<close> paper\<close>
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    35
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    36
lemma "(\<Union>x \<in> C. f x \<union> g x) = \<Union>(f ` C)  \<union>  \<Union>(g ` C)"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    37
  \<comment> \<open>Union-image, called \<open>Un_Union_image\<close> in Main HOL\<close>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    38
  by blast
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    39
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    40
lemma "(\<Inter>x \<in> C. f x \<inter> g x) = \<Inter>(f ` C) \<inter> \<Inter>(g ` C)"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    41
  \<comment> \<open>Inter-image, called \<open>Int_Inter_image\<close> in Main HOL\<close>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    42
  by blast
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    43
16898
543ee8fabe1a revised examples
paulson
parents: 16417
diff changeset
    44
lemma singleton_example_1:
543ee8fabe1a revised examples
paulson
parents: 16417
diff changeset
    45
     "\<And>S::'a set set. \<forall>x \<in> S. \<forall>y \<in> S. x \<subseteq> y \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
18391
2e901da7cd3a meson no longer does these examples
paulson
parents: 16898
diff changeset
    46
  by blast
16898
543ee8fabe1a revised examples
paulson
parents: 16417
diff changeset
    47
543ee8fabe1a revised examples
paulson
parents: 16417
diff changeset
    48
lemma singleton_example_2:
543ee8fabe1a revised examples
paulson
parents: 16417
diff changeset
    49
     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    50
  \<comment> \<open>Variant of the problem above.\<close>
18391
2e901da7cd3a meson no longer does these examples
paulson
parents: 16898
diff changeset
    51
  by blast
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    52
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    53
lemma "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    54
  \<comment> \<open>A unique fixpoint theorem --- \<open>fast\<close>/\<open>best\<close>/\<open>meson\<close> all fail.\<close>
24573
paulson
parents: 19982
diff changeset
    55
  by metis
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    56
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    57
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
    58
subsection \<open>Cantor's Theorem: There is no surjection from a set to its powerset\<close>
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    59
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    60
lemma cantor1: "\<not> (\<exists>f:: 'a \<Rightarrow> 'a set. \<forall>S. \<exists>x. f x = S)"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    61
  \<comment> \<open>Requires best-first search because it is undirectional.\<close>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    62
  by best
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    63
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
    64
schematic_goal "\<forall>f:: 'a \<Rightarrow> 'a set. \<forall>x. f x \<noteq> ?S f"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    65
  \<comment> \<open>This form displays the diagonal term.\<close>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    66
  by best
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    67
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
    68
schematic_goal "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    69
  \<comment> \<open>This form exploits the set constructs.\<close>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    70
  by (rule notI, erule rangeE, best)
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    71
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
    72
schematic_goal "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
    73
  \<comment> \<open>Or just this!\<close>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    74
  by best
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    75
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    76
61937
wenzelm
parents: 61933
diff changeset
    77
subsection \<open>The Schröder-Bernstein Theorem\<close>
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    78
61937
wenzelm
parents: 61933
diff changeset
    79
lemma disj_lemma: "- (f ` X) = g' ` (-X) \<Longrightarrow> f a = g' b \<Longrightarrow> a \<in> X \<Longrightarrow> b \<in> X"
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    80
  by blast
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    81
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    82
lemma surj_if_then_else:
61937
wenzelm
parents: 61933
diff changeset
    83
  "-(f ` X) = g' ` (-X) \<Longrightarrow> surj (\<lambda>z. if z \<in> X then f z else g' z)"
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    84
  by (simp add: surj_def) blast
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    85
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    86
lemma bij_if_then_else:
61937
wenzelm
parents: 61933
diff changeset
    87
  "inj_on f X \<Longrightarrow> inj_on g' (-X) \<Longrightarrow> -(f ` X) = g' ` (-X) \<Longrightarrow>
wenzelm
parents: 61933
diff changeset
    88
    h = (\<lambda>z. if z \<in> X then f z else g' z) \<Longrightarrow> inj h \<and> surj h"
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    89
  apply (unfold inj_on_def)
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    90
  apply (simp add: surj_if_then_else)
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    91
  apply (blast dest: disj_lemma sym)
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    92
  done
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    93
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    94
lemma decomposition: "\<exists>X. X = - (g ` (- (f ` X)))"
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    95
  apply (rule exI)
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    96
  apply (rule lfp_unfold)
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    97
  apply (rule monoI, blast)
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
    98
  done
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
    99
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   100
theorem Schroeder_Bernstein:
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   101
  "inj (f :: 'a \<Rightarrow> 'b) \<Longrightarrow> inj (g :: 'b \<Rightarrow> 'a)
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   102
    \<Longrightarrow> \<exists>h:: 'a \<Rightarrow> 'b. inj h \<and> surj h"
15488
7c638a46dcbb tidying of some subst/simplesubst proofs
paulson
parents: 15481
diff changeset
   103
  apply (rule decomposition [where f=f and g=g, THEN exE])
7c638a46dcbb tidying of some subst/simplesubst proofs
paulson
parents: 15481
diff changeset
   104
  apply (rule_tac x = "(\<lambda>z. if z \<in> x then f z else inv g z)" in exI) 
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   105
    \<comment>\<open>The term above can be synthesized by a sufficiently detailed proof.\<close>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   106
  apply (rule bij_if_then_else)
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   107
     apply (rule_tac [4] refl)
33057
764547b68538 inv_onto -> inv_into
nipkow
parents: 32988
diff changeset
   108
    apply (rule_tac [2] inj_on_inv_into)
15306
51f3d31e8eea fixed proof
nipkow
parents: 14353
diff changeset
   109
    apply (erule subset_inj_on [OF _ subset_UNIV])
15488
7c638a46dcbb tidying of some subst/simplesubst proofs
paulson
parents: 15481
diff changeset
   110
   apply blast
63804
70554522bf98 standardized alias;
wenzelm
parents: 61945
diff changeset
   111
  apply (erule ssubst, subst double_complement, erule image_inv_f_f [symmetric])
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   112
  done
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
   113
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
   114
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   115
subsection \<open>A simple party theorem\<close>
24853
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   116
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   117
text\<open>\emph{At any party there are two people who know the same
24853
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   118
number of people}. Provided the party consists of at least two people
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   119
and the knows relation is symmetric. Knowing yourself does not count
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   120
--- otherwise knows needs to be reflexive. (From Freek Wiedijk's talk
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   121
at TPHOLs 2007.)\<close>
24853
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   122
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   123
lemma equal_number_of_acquaintances:
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   124
assumes "Domain R <= A" and "sym R" and "card A \<ge> 2"
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   125
shows "\<not> inj_on (%a. card(R `` {a} - {a})) A"
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   126
proof -
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   127
  let ?N = "%a. card(R `` {a} - {a})"
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   128
  let ?n = "card A"
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   129
  have "finite A" using \<open>card A \<ge> 2\<close> by(auto intro:ccontr)
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   130
  have 0: "R `` A <= A" using \<open>sym R\<close> \<open>Domain R <= A\<close>
46752
e9e7209eb375 more fundamental pred-to-set conversions, particularly by means of inductive_set; associated consolidation of some theorem names (c.f. NEWS)
haftmann
parents: 45966
diff changeset
   131
    unfolding Domain_unfold sym_def by blast
24853
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   132
  have h: "ALL a:A. R `` {a} <= A" using 0 by blast
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   133
  hence 1: "ALL a:A. finite(R `` {a})" using \<open>finite A\<close>
24853
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   134
    by(blast intro: finite_subset)
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   135
  have sub: "?N ` A <= {0..<?n}"
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   136
  proof -
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   137
    have "ALL a:A. R `` {a} - {a} < A" using h by blast
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   138
    thus ?thesis using psubset_card_mono[OF \<open>finite A\<close>] by auto
24853
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   139
  qed
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   140
  show "~ inj_on ?N A" (is "~ ?I")
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   141
  proof
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   142
    assume ?I
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   143
    hence "?n = card(?N ` A)" by(rule card_image[symmetric])
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   144
    with sub \<open>finite A\<close> have 2[simp]: "?N ` A = {0..<?n}"
24853
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   145
      using subset_card_intvl_is_intvl[of _ 0] by(auto)
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   146
    have "0 : ?N ` A" and "?n - 1 : ?N ` A"  using \<open>card A \<ge> 2\<close> by simp+
24853
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   147
    then obtain a b where ab: "a:A" "b:A" and Na: "?N a = 0" and Nb: "?N b = ?n - 1"
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   148
      by (auto simp del: 2)
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   149
    have "a \<noteq> b" using Na Nb \<open>card A \<ge> 2\<close> by auto
24853
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   150
    have "R `` {a} - {a} = {}" by (metis 1 Na ab card_eq_0_iff finite_Diff)
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   151
    hence "b \<notin> R `` {a}" using \<open>a\<noteq>b\<close> by blast
24853
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   152
    hence "a \<notin> R `` {b}" by (metis Image_singleton_iff assms(2) sym_def)
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   153
    hence 3: "R `` {b} - {b} <= A - {a,b}" using 0 ab by blast
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   154
    have 4: "finite (A - {a,b})" using \<open>finite A\<close> by simp
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   155
    have "?N b <= ?n - 2" using ab \<open>a\<noteq>b\<close> \<open>finite A\<close> card_mono[OF 4 3] by simp
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   156
    then show False using Nb \<open>card A \<ge>  2\<close> by arith
24853
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   157
  qed
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   158
qed
aab5798e5a33 added lemmas
nipkow
parents: 24573
diff changeset
   159
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   160
text \<open>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   161
  From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   162
  293-314.
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   163
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   164
  Isabelle can prove the easy examples without any special mechanisms,
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   165
  but it can't prove the hard ones.
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   166
\<close>
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
   167
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   168
lemma "\<exists>A. (\<forall>x \<in> A. x \<le> (0::int))"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   169
  \<comment> \<open>Example 1, page 295.\<close>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   170
  by force
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
   171
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   172
lemma "D \<in> F \<Longrightarrow> \<exists>G. \<forall>A \<in> G. \<exists>B \<in> F. A \<subseteq> B"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   173
  \<comment> \<open>Example 2.\<close>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   174
  by force
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
   175
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   176
lemma "P a \<Longrightarrow> \<exists>A. (\<forall>x \<in> A. P x) \<and> (\<exists>y. y \<in> A)"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   177
  \<comment> \<open>Example 3.\<close>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   178
  by force
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
   179
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   180
lemma "a < b \<and> b < (c::int) \<Longrightarrow> \<exists>A. a \<notin> A \<and> b \<in> A \<and> c \<notin> A"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   181
  \<comment> \<open>Example 4.\<close>
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   182
  by auto \<comment>\<open>slow\<close>
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
   183
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   184
lemma "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   185
  \<comment> \<open>Example 5, page 298.\<close>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   186
  by force
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
   187
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   188
lemma "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   189
  \<comment> \<open>Example 6.\<close>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   190
  by force
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
   191
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   192
lemma "\<exists>A. a \<notin> A"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   193
  \<comment> \<open>Example 7.\<close>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   194
  by force
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
   195
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61937
diff changeset
   196
lemma "(\<forall>u v. u < (0::int) \<longrightarrow> u \<noteq> \<bar>v\<bar>)
1135b8de26c3 more symbols;
wenzelm
parents: 61937
diff changeset
   197
    \<longrightarrow> (\<exists>A::int set. -2 \<in> A & (\<forall>y. \<bar>y\<bar> \<notin> A))"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   198
  \<comment> \<open>Example 8 needs a small hint.\<close>
45966
03ce2b2a29a2 tuned proofs
haftmann
parents: 44276
diff changeset
   199
  by force
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   200
    \<comment> \<open>not \<open>blast\<close>, which can't simplify \<open>-2 < 0\<close>\<close>
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
   201
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   202
text \<open>Example 9 omitted (requires the reals).\<close>
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
   203
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   204
text \<open>The paper has no Example 10!\<close>
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
   205
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   206
lemma "(\<forall>A. 0 \<in> A \<and> (\<forall>x \<in> A. Suc x \<in> A) \<longrightarrow> n \<in> A) \<and>
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   207
  P 0 \<and> (\<forall>x. P x \<longrightarrow> P (Suc x)) \<longrightarrow> P n"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   208
  \<comment> \<open>Example 11: needs a hint.\<close>
40928
ace26e2cee91 eliminated unqualified accesses of datatype facts -- it seems like they all of them were unintended
krauss
parents: 36319
diff changeset
   209
by(metis nat.induct)
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
   210
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   211
lemma
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   212
  "(\<forall>A. (0, 0) \<in> A \<and> (\<forall>x y. (x, y) \<in> A \<longrightarrow> (Suc x, Suc y) \<in> A) \<longrightarrow> (n, m) \<in> A)
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   213
    \<and> P n \<longrightarrow> P m"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   214
  \<comment> \<open>Example 12.\<close>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   215
  by auto
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
   216
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   217
lemma
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   218
  "(\<forall>x. (\<exists>u. x = 2 * u) = (\<not> (\<exists>v. Suc x = 2 * v))) \<longrightarrow>
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   219
    (\<exists>A. \<forall>x. (x \<in> A) = (Suc x \<notin> A))"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 61343
diff changeset
   220
  \<comment> \<open>Example EO1: typo in article, and with the obvious fix it seems
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 61337
diff changeset
   221
      to require arithmetic reasoning.\<close>
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   222
  apply clarify
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   223
  apply (rule_tac x = "{x. \<exists>u. x = 2 * u}" in exI, auto)
34055
fdf294ee08b2 streamlined proofs
paulson
parents: 33057
diff changeset
   224
   apply metis+
13107
8743cc847224 tuned presentation;
wenzelm
parents: 13058
diff changeset
   225
  done
13058
ad6106d7b4bb converted theory "set" to Isar and added some SET-VAR examples
paulson
parents: 9100
diff changeset
   226
9100
9e081c812338 fixed deps;
wenzelm
parents:
diff changeset
   227
end