[1062] | 1 | include "basics/jmeq.ma". |
---|
| 2 | include "basics/types.ma". |
---|
| 3 | include "basics/list.ma". |
---|
| 4 | |
---|
| 5 | notation > "hvbox(a break ≃ b)" |
---|
| 6 | non associative with precedence 45 |
---|
| 7 | for @{ 'jmeq ? $a ? $b }. |
---|
| 8 | |
---|
| 9 | notation < "hvbox(term 46 a break maction (≃) (≃\sub(t,u)) term 46 b)" |
---|
| 10 | non associative with precedence 45 |
---|
| 11 | for @{ 'jmeq $t $a $u $b }. |
---|
| 12 | |
---|
| 13 | interpretation "john major's equality" 'jmeq t x u y = (jmeq t x u y). |
---|
| 14 | |
---|
| 15 | lemma eq_to_jmeq: |
---|
| 16 | ∀A: Type[0]. |
---|
| 17 | ∀x, y: A. |
---|
| 18 | x = y → x ≃ y. |
---|
| 19 | // |
---|
| 20 | qed. |
---|
| 21 | |
---|
| 22 | definition inject : ∀A.∀P:A → Prop.∀a.∀p:P a.Σx:A.P x ≝ λA,P,a,p. dp … a p. |
---|
| 23 | definition eject : ∀A.∀P: A → Prop.(Σx:A.P x) → A ≝ λA,P,c.match c with [ dp w p ⇒ w]. |
---|
| 24 | |
---|
| 25 | coercion inject nocomposites: ∀A.∀P:A → Prop.∀a.∀p:P a.Σx:A.P x ≝ inject on a:? to Σx:?.?. |
---|
| 26 | coercion eject nocomposites: ∀A.∀P:A → Prop.∀c:Σx:A.P x.A ≝ eject on _c:Σx:?.? to ?. |
---|
| 27 | |
---|
| 28 | (*axiom VOID: Type[0]. |
---|
| 29 | axiom assert_false: VOID. |
---|
| 30 | definition bigbang: ∀A:Type[0].False → VOID → A. |
---|
| 31 | #A #abs cases abs |
---|
| 32 | qed. |
---|
| 33 | |
---|
| 34 | coercion bigbang nocomposites: ∀A:Type[0].False → ∀v:VOID.A ≝ bigbang on _v:VOID to ?.*) |
---|
| 35 | |
---|
| 36 | lemma sig2: ∀A.∀P:A → Prop. ∀p:Σx:A.P x. P (eject … p). |
---|
| 37 | #A #P #p cases p #w #q @q |
---|
| 38 | qed. |
---|
| 39 | |
---|
| 40 | lemma jmeq_to_eq: ∀A:Type[0]. ∀x,y:A. x≃y → x=y. |
---|
| 41 | #A #x #y #JMEQ @(jmeq_elim ? x … JMEQ) % |
---|
| 42 | qed. |
---|
| 43 | |
---|
| 44 | coercion jmeq_to_eq: ∀A:Type[0]. ∀x,y:A. ∀p:x≃y.x=y ≝ jmeq_to_eq on _p:?≃? to ?=?. |
---|
| 45 | |
---|
| 46 | (* END RUSSELL **) |
---|