include "basics/jmeq.ma". include "basics/types.ma". include "basics/list.ma". notation > "hvbox(a break ≃ b)" non associative with precedence 45 for @{ 'jmeq ? $a ? $b }. notation < "hvbox(term 46 a break maction (≃) (≃\sub(t,u)) term 46 b)" non associative with precedence 45 for @{ 'jmeq $t $a $u $b }. interpretation "john major's equality" 'jmeq t x u y = (jmeq t x u y). lemma eq_to_jmeq: ∀A: Type[0]. ∀x, y: A. x = y → x ≃ y. // qed. 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. lemma jmeq_to_eq: ∀A:Type[0]. ∀x,y:A. x≃y → x=y. #A #x #y #JMEQ @(jmeq_elim ? x … JMEQ) % qed. coercion jmeq_to_eq: ∀A:Type[0]. ∀x,y:A. ∀p:x≃y.x=y ≝ jmeq_to_eq on _p:?≃? to ?=?. (* END RUSSELL **)