From ac1263c3471b351f520d7e2abf217220efc10e7e Mon Sep 17 00:00:00 2001 From: Claudio Sacerdoti Coen Date: Sun, 28 Dec 2008 16:52:37 +0000 Subject: [PATCH 1/1] Many universe inconsistency avoided here and there. --- .../formal_topology/overlap/subsets.ma | 33 ++++++++++++------- 1 file changed, 22 insertions(+), 11 deletions(-) diff --git a/helm/software/matita/contribs/formal_topology/overlap/subsets.ma b/helm/software/matita/contribs/formal_topology/overlap/subsets.ma index 4bbe7fdbe..9847d4c55 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/subsets.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/subsets.ma @@ -153,24 +153,26 @@ definition big_intersects: | apply (. (#‡(e i)\sup -1)); apply f]] qed. -axiom (*definition*) big_union: +(* senza questo exT "fresco", universe inconsistency *) +inductive exT (A:Type) (P:A→CProp) : CProp ≝ + ex_introT: ∀w:A. P w → exT A P. + +definition big_union: ∀A:SET.∀I:SET.unary_morphism2 (setoid1_of_setoid I ⇒ Ω \sup A) (setoid2_of_setoid1 (Ω \sup A)). -(* intros; constructor 1; + intros; constructor 1; [ intro; whd; whd in A; whd in I; - apply ({x | ∃i:I. x ∈ t i}); + apply ({x | (*∃i:carr I. x ∈ t i*) exT (carr I) (λi. x ∈ t i)}); simplify; intros; split; intros; cases H; clear H; exists; [1,3:apply w] [ apply (. (e‡#)); | apply (. (e \sup -1‡#)); ] apply x; | intros; split; intros 2; simplify in f ⊢ %; cases f; clear f; exists; [1,3:apply w] [ apply (. (#‡(e w))); apply x; | apply (. (#‡(e w)\sup -1)); apply x]] -qed.*) +qed. (* incluso prima non funziona piu' nulla *) include "o-algebra.ma". - -axiom daemon: False. definition SUBSETS: objs1 SET → OAlgebra. intro A; constructor 1; [ apply (Ω \sup A); @@ -187,18 +189,27 @@ definition SUBSETS: objs1 SET → OAlgebra. | intros; apply transitive_subseteq_operator; [2: apply f; | skip | assumption] | intros; cases f; exists [apply w] assumption | intros; intros 2; apply (f ? f1 i); - |(* intros; intros 2; apply f; exists; [apply i] assumption;*) + | intros; intros 2; apply f; + (* senza questa change, universe inconsistency *) + whd; change in ⊢ (? ? (λ_:%.?)) with (carr I); + exists; [apply i] assumption; | intros 3; cases f; | intros 3; constructor 1; | intros; cases f; exists; [apply w] [ assumption | whd; intros; cases i; simplify; assumption] - |(* intros; split; intro; - [ cases f; cases x1; exists [apply w1] exists [apply w] assumption; - | cases H; cases x; exists; [apply w1] [assumption | exists; [apply w] assumption]]*) + | intros; split; intro; + [ cases f; cases x1; + (* senza questa change, universe inconsistency *) + change in ⊢ (? ? (λ_:%.?)) with (carr I); + exists [apply w1] exists [apply w] assumption; + | cases H; cases x; exists; [apply w1] + [ assumption + | (* senza questa change, universe inconsistency *) + whd; change in ⊢ (? ? (λ_:%.?)) with (carr I); + exists; [apply w] assumption]] | intros; intros 2; cases (H (singleton ? a) ?); [ exists; [apply a] [assumption | change with (a = a); apply refl1;] | change in x1 with (a = w); change with (mem A a q); apply (. (x1 \sup -1‡#)); assumption]] -cases daemon; qed. \ No newline at end of file -- 2.39.2