1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "demo/formal_topology.ma".
16 include "datatypes/constructors.ma".
19 axiom S: A → Ω \sup A.
21 definition axs: axiom_set.
25 | intros; apply (S a)]
28 definition S' : axs → Ω \sup axs ≝ S.
30 definition emptyset: Ω \sup axs ≝ {x | False}.
32 definition Z: Ω \sup axs ≝ {x | x ◃ emptyset}.
34 theorem cantor: ∀a:axs. ¬ (Z ⊆ S' a ∧ S' a ⊆ Z).
35 intros 2; cases H; clear H;
37 [2: constructor 2; [constructor 1 | whd in ⊢ (? ? ? % ?); autobatch]]
39 [2: apply transitivity; [apply (S' a)]
41 | constructor 1; intros;
42 lapply (H2 ? H); whd in Hletin; assumption]]
44 [2: apply H1; whd; assumption]
45 generalize in match Hcut2; clear Hcut2;
46 change with (a ∈ {a | a ∈ S' a → False});
47 apply (covers_elim ?????? Hcut1);
48 [ intros 2; simplify; intros; assumption;
49 | intros; simplify; intros; whd in H3; simplify in H3; apply (H3 a1);