src/HOL/Metis_Examples/Sets.thy
author haftmann
Sat, 24 Dec 2011 15:53:11 +0100
changeset 45972 deda685ba210
parent 43197 c71657bbdbc0
child 46077 86e6e9d42ad7
permissions -rw-r--r--
commented out examples which choke on strict set/pred distinction
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43197
c71657bbdbc0 tuned Metis examples
blanchet
parents: 42103
diff changeset
     1
(*  Title:      HOL/Metis_Examples/Sets.thy
c71657bbdbc0 tuned Metis examples
blanchet
parents: 42103
diff changeset
     2
    Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
41144
509e51b7509a example tuning
blanchet
parents: 37325
diff changeset
     3
    Author:     Jasmin Blanchette, TU Muenchen
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
     4
43197
c71657bbdbc0 tuned Metis examples
blanchet
parents: 42103
diff changeset
     5
Metis example featuring typed set theory.
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
     6
*)
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
     7
43197
c71657bbdbc0 tuned Metis examples
blanchet
parents: 42103
diff changeset
     8
header {* Metis Example Featuring Typed Set Theory *}
c71657bbdbc0 tuned Metis examples
blanchet
parents: 42103
diff changeset
     9
c71657bbdbc0 tuned Metis examples
blanchet
parents: 42103
diff changeset
    10
theory Sets
33027
9cf389429f6d modernized session Metis_Examples;
wenzelm
parents: 32864
diff changeset
    11
imports Main
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
    12
begin
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
    13
42103
6066a35f6678 Metis examples use the new Skolemizer to test it
blanchet
parents: 41144
diff changeset
    14
declare [[metis_new_skolemizer]]
6066a35f6678 Metis examples use the new Skolemizer to test it
blanchet
parents: 41144
diff changeset
    15
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
    16
lemma "EX x X. ALL y. EX z Z. (~P(y,y) | P(x,x) | ~S(z,x)) &
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
    17
               (S(x,y) | ~S(y,z) | Q(Z,Z))  &
43197
c71657bbdbc0 tuned Metis examples
blanchet
parents: 42103
diff changeset
    18
               (Q(X,y) | ~Q(y,Z) | S(X,X))"
23519
a4ffa756d8eb bug fixes to proof reconstruction
paulson
parents: 23449
diff changeset
    19
by metis
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
    20
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
    21
lemma "P(n::nat) ==> ~P(0) ==> n ~= 0"
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
    22
by metis
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
    23
36925
ffad77bb3046 renamed Sledgehammer options
blanchet
parents: 36573
diff changeset
    24
sledgehammer_params [isar_proof, isar_shrink_factor = 1]
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
    25
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
    26
(*multiple versions of this example*)
45972
deda685ba210 commented out examples which choke on strict set/pred distinction
haftmann
parents: 43197
diff changeset
    27
(* lemma (*equal_union: *)
36566
f91342f218a9 redid some Sledgehammer/Metis proofs
blanchet
parents: 36407
diff changeset
    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: 36925
diff changeset
    29
proof -
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff changeset
    30
  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
    31
  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
    32
  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
    33
  { 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
    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 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
    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 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
    37
    { 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
    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 F1) }
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 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
    41
      { 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
    42
        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
    43
      moreover
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff 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: 36925
diff 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: 36925
diff changeset
    46
        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
    47
        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
    48
        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
    49
      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
    50
    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
    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 "\<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
    53
    { 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
    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 F1) }
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 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
    57
      { 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
    58
        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
    59
      moreover
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff 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: 36925
diff 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: 36925
diff changeset
    62
        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
    63
        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
    64
        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
    65
      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
    66
    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
    67
  moreover
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff changeset
    68
  { 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
    69
    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
    70
  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
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
    71
qed
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
    72
36925
ffad77bb3046 renamed Sledgehammer options
blanchet
parents: 36573
diff changeset
    73
sledgehammer_params [isar_proof, isar_shrink_factor = 2]
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
    74
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
    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: 36925
diff 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: 36925
diff changeset
    77
proof -
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff changeset
    78
  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
    79
  { 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
    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: 36925
diff changeset
    81
      { 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
    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
      moreover
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff changeset
    84
      { 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
    85
        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
    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_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 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
    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> 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
    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_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
    91
    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
    92
  moreover
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff changeset
    93
  { 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
    94
    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
    95
  moreover
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> 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
    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_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
    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 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
   100
    { 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
   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
    moreover
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff changeset
   103
    { 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
   104
      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
   105
      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
   106
    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
   107
  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
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   108
qed
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   109
36925
ffad77bb3046 renamed Sledgehammer options
blanchet
parents: 36573
diff changeset
   110
sledgehammer_params [isar_proof, isar_shrink_factor = 3]
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   111
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   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: 36925
diff 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: 36925
diff changeset
   114
proof -
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff changeset
   115
  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
   116
  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
   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
  moreover
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff changeset
   120
  { 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
   121
    { 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
   122
      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
   123
    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
   124
  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
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   125
qed
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   126
36925
ffad77bb3046 renamed Sledgehammer options
blanchet
parents: 36573
diff changeset
   127
sledgehammer_params [isar_proof, isar_shrink_factor = 4]
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   128
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   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: 36925
diff 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: 36925
diff changeset
   131
proof -
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff changeset
   132
  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
   133
  { 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
   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_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
   135
  moreover
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff changeset
   136
  { 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
   137
    { 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
   138
      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
   139
    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
   140
  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
   141
qed
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   142
37325
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff changeset
   143
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
   144
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   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: 36925
diff 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
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   147
by (metis Un_least Un_upper1 Un_upper2 set_eq_subset)
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   148
37325
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff 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: 36925
diff changeset
   150
by (metis Int_greatest Int_lower1 Int_lower2 subset_antisym)
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   151
37325
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff changeset
   152
lemma fixedpoint: "\<exists>!x. f (g x) = x \<Longrightarrow> \<exists>!y. g (f y) = y"
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   153
by metis
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   154
37325
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff 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: 36925
diff changeset
   156
proof -
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff 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: 36925
diff changeset
   158
  thus "\<exists>!y\<Colon>'b. g (f y) = y" by metis
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   159
qed
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   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
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   162
     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   163
by (metis Set.subsetI Union_upper insertCI set_eq_subset)
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   164
37325
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff changeset
   165
lemma (* singleton_example_2: *)
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   166
     "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
32685
29e4e567b5f4 tuned proofs
haftmann
parents: 32519
diff changeset
   167
by (metis Set.subsetI Union_upper insert_iff set_eq_subset)
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   168
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   169
lemma singleton_example_2:
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   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: 36925
diff changeset
   171
proof -
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff changeset
   172
  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
   173
  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
   174
  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
   175
  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
   176
  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
   177
  thus "\<exists>z. S \<subseteq> {z}" by metis
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   178
qed
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   179
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   180
text {*
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   181
  From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   182
  293-314.
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   183
*}
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   184
37325
c2a44bc874f9 redid the Isar proofs using the latest Sledgehammer, eliminating the last occurrences of "neg_clausify" in proofs
blanchet
parents: 36925
diff 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: 36925
diff changeset
   186
   (2) We must rename set variables to avoid type clashes. *)
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   187
lemma "\<exists>B. (\<forall>x \<in> B. x \<le> (0::int))"
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   188
      "D \<in> F \<Longrightarrow> \<exists>G. \<forall>A \<in> G. \<exists>B \<in> F. A \<subseteq> B"
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   189
      "P a \<Longrightarrow> \<exists>A. (\<forall>x \<in> A. P x) \<and> (\<exists>y. y \<in> A)"
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   190
      "a < b \<and> b < (c::int) \<Longrightarrow> \<exists>B. a \<notin> B \<and> b \<in> B \<and> c \<notin> B"
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   191
      "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   192
      "P (f b) \<Longrightarrow> \<exists>s A. (\<forall>x \<in> A. P x) \<and> f s \<in> A"
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   193
      "\<exists>A. a \<notin> A"
36566
f91342f218a9 redid some Sledgehammer/Metis proofs
blanchet
parents: 36407
diff changeset
   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: 36925
diff 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: 36925
diff changeset
   196
      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
   197
     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
   198
    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
   199
   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
   200
  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
   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: 36925
diff changeset
   202
by (metis pair_in_Id_conv)
45972
deda685ba210 commented out examples which choke on strict set/pred distinction
haftmann
parents: 43197
diff changeset
   203
*)
23449
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   204
dd874e6a3282 integration of Metis prover
paulson
parents:
diff changeset
   205
end