]> matita.cs.unibo.it Git - helm.git/blob - weblib/Cerco/ASM/JMCoercions.ma
3f43e6b706080c8a7ce31c739fa393698743c172
[helm.git] / weblib / Cerco / ASM / JMCoercions.ma
1 include "basics/jmeq.ma".
2 include "basics/types.ma".
3 include "basics/list.ma".
4
5 definition inject : ∀A.∀P:A → Prop.∀a.∀p:P a.Σx:A.P x ≝ λA,P,a,p. dp … a p.
6 definition eject : ∀A.∀P: A → Prop.(Σx:A.P x) → A ≝ λA,P,c.match c with [ dp w p ⇒ w].
7
8 coercion inject nocomposites: ∀A.∀P:A → Prop.∀a.∀p:P a.Σx:A.P x ≝ inject on a:? to Σx:?.?.
9 coercion eject nocomposites: ∀A.∀P:A → Prop.∀c:Σx:A.P x.A ≝ eject on _c:Σx:?.? to ?.
10
11 (*axiom VOID: Type[0].
12 axiom assert_false: VOID.
13 definition bigbang: ∀A:Type[0].False → VOID → A.
14  #A #abs cases abs
15 qed.
16
17 coercion bigbang nocomposites: ∀A:Type[0].False → ∀v:VOID.A ≝ bigbang on _v:VOID to ?.*)
18
19 lemma sig2: ∀A.∀P:A → Prop. ∀p:Σx:A.P x. P (eject … p).
20  #A #P #p cases p #w #q @q
21 qed.
22
23 (* END RUSSELL **)