X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=weblib%2FCerco%2FASM%2FJMCoercions.ma;h=f36d310c08c1cbcf009df9c9e74a544141d83240;hb=50a9ed8c6207145fccf59e6a5dbbff935cd2c6d7;hp=3f43e6b706080c8a7ce31c739fa393698743c172;hpb=b5f54d2815f446a999736abd0ffe80641596a5f6;p=helm.git diff --git a/weblib/Cerco/ASM/JMCoercions.ma b/weblib/Cerco/ASM/JMCoercions.ma index 3f43e6b70..f36d310c0 100644 --- a/weblib/Cerco/ASM/JMCoercions.ma +++ b/weblib/Cerco/ASM/JMCoercions.ma @@ -2,11 +2,11 @@ include "basics/jmeq.ma". include "basics/types.ma". include "basics/list.ma". -definition inject : ∀A.∀P:A → Prop.∀a.∀p:P a.Σx:A.P x ≝ λA,P,a,p. dp … a p. -definition eject : ∀A.∀P: A → Prop.(Σx:A.P x) → A ≝ λA,P,c.match c with [ dp w p ⇒ w]. +definition inject : ∀A.∀P:A → Prop.∀a.∀p:P a.a title="Sigma" href="cic:/fakeuri.def(1)"Σ/ax:A.P x ≝ λA,P,a,p. a href="cic:/matita/basics/types/Sig.con(0,1,2)"dp/a … a p. +definition eject : ∀A.∀P: A → Prop.(a title="Sigma" href="cic:/fakeuri.def(1)"Σ/ax:A.P x) → A ≝ λA,P,c.match c with [ dp w p ⇒ w]. -coercion inject nocomposites: ∀A.∀P:A → Prop.∀a.∀p:P a.Σx:A.P x ≝ inject on a:? to Σx:?.?. -coercion eject nocomposites: ∀A.∀P:A → Prop.∀c:Σx:A.P x.A ≝ eject on _c:Σx:?.? to ?. +(* coercion inject nocomposites: ∀A.∀P:A → Prop.∀a.∀p:P a.Σx:A.P x ≝ inject on a:? to Σx:?.?. +coercion eject nocomposites: ∀A.∀P:A → Prop.∀c:Σx:A.P x.A ≝ eject on _c:Σx:?.? to ?. *) (*axiom VOID: Type[0]. axiom assert_false: VOID. @@ -16,8 +16,8 @@ qed. coercion bigbang nocomposites: ∀A:Type[0].False → ∀v:VOID.A ≝ bigbang on _v:VOID to ?.*) -lemma sig2: ∀A.∀P:A → Prop. ∀p:Σx:A.P x. P (eject … p). +lemma sig2: ∀A.∀P:A → Prop. ∀p:a title="Sigma" href="cic:/fakeuri.def(1)"Σ/ax:A.P x. P (a href="cic:/matita/Cerco/ASM/JMCoercions/eject.def(1)"eject/a … p). #A #P #p cases p #w #q @q qed. -(* END RUSSELL **) +(* END RUSSELL **) \ No newline at end of file