X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=weblib%2FCerco%2FASM%2FJMCoercions.ma;fp=weblib%2FCerco%2FASM%2FJMCoercions.ma;h=3f43e6b706080c8a7ce31c739fa393698743c172;hb=b5f54d2815f446a999736abd0ffe80641596a5f6;hp=0000000000000000000000000000000000000000;hpb=704bbf749f44396c1a610f336b4e1cf0d25e9370;p=helm.git diff --git a/weblib/Cerco/ASM/JMCoercions.ma b/weblib/Cerco/ASM/JMCoercions.ma new file mode 100644 index 000000000..3f43e6b70 --- /dev/null +++ b/weblib/Cerco/ASM/JMCoercions.ma @@ -0,0 +1,23 @@ +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]. + +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. +definition bigbang: ∀A:Type[0].False → VOID → A. + #A #abs cases abs +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). + #A #P #p cases p #w #q @q +qed. + +(* END RUSSELL **)