1 | (* RUSSEL **) |
---|
2 | |
---|
3 | include "basics/jmeq.ma". |
---|
4 | include "basics/types.ma". |
---|
5 | include "basics/list.ma". |
---|
6 | |
---|
7 | notation > "hvbox(a break ≃ b)" |
---|
8 | non associative with precedence 45 |
---|
9 | for @{ 'jmeq ? $a ? $b }. |
---|
10 | |
---|
11 | notation < "hvbox(term 46 a break maction (≃) (≃\sub(t,u)) term 46 b)" |
---|
12 | non associative with precedence 45 |
---|
13 | for @{ 'jmeq $t $a $u $b }. |
---|
14 | |
---|
15 | interpretation "john major's equality" 'jmeq t x u y = (jmeq t x u y). |
---|
16 | |
---|
17 | lemma eq_to_jmeq: |
---|
18 | ∀A: Type[0]. |
---|
19 | ∀x, y: A. |
---|
20 | x = y → x ≃ y. |
---|
21 | // |
---|
22 | qed. |
---|
23 | |
---|
24 | definition inject : ∀A.∀P:A → Prop.∀a.∀p:P a.Σx:A.P x ≝ λA,P,a,p. dp … a p. |
---|
25 | definition eject : ∀A.∀P: A → Prop.(Σx:A.P x) → A ≝ λA,P,c.match c with [ dp w p ⇒ w]. |
---|
26 | |
---|
27 | coercion inject nocomposites: ∀A.∀P:A → Prop.∀a.∀p:P a.Σx:A.P x ≝ inject on a:? to Σx:?.?. |
---|
28 | coercion eject nocomposites: ∀A.∀P:A → Prop.∀c:Σx:A.P x.A ≝ eject on _c:Σx:?.? to ?. |
---|
29 | |
---|
30 | axiom VOID: Type[0]. |
---|
31 | axiom assert_false: VOID. |
---|
32 | definition bigbang: ∀A:Type[0].False → VOID → A. |
---|
33 | #A #abs cases abs |
---|
34 | qed. |
---|
35 | |
---|
36 | coercion bigbang nocomposites: ∀A:Type[0].False → ∀v:VOID.A ≝ bigbang on _v:VOID to ?. |
---|
37 | |
---|
38 | lemma sig2: ∀A.∀P:A → Prop. ∀p:Σx:A.P x. P (eject … p). |
---|
39 | #A #P #p cases p #w #q @q |
---|
40 | qed. |
---|
41 | |
---|
42 | lemma jmeq_to_eq: ∀A:Type[0]. ∀x,y:A. x≃y → x=y. |
---|
43 | #A #x #y #JMEQ @(jmeq_elim ? x … JMEQ) % |
---|
44 | qed. |
---|
45 | |
---|
46 | coercion jmeq_to_eq: ∀A:Type[0]. ∀x,y:A. ∀p:x≃y.x=y ≝ jmeq_to_eq on _p:?≃? to ?=?. |
---|
47 | |
---|
48 | (* END RUSSELL **) |
---|
49 | |
---|
50 | include "ASM/Util.ma". |
---|
51 | |
---|
52 | let rec foldl_strong_internal |
---|
53 | (A: Type[0]) (P: list A → Type[0]) (l: list A) |
---|
54 | (H: ∀prefix. ∀hd. ∀tl. l = prefix @ [hd] @ tl → P prefix → P (prefix @ [hd])) |
---|
55 | (prefix: list A) (suffix: list A) (acc: P prefix) on suffix: |
---|
56 | l = prefix @ suffix → P(prefix @ suffix) ≝ |
---|
57 | match suffix return λl'. l = prefix @ l' → P (prefix @ l') with |
---|
58 | [ nil ⇒ λprf. ? |
---|
59 | | cons hd tl ⇒ λprf. ? |
---|
60 | ]. |
---|
61 | [ > (append_nil ?) |
---|
62 | @ acc |
---|
63 | | applyS (foldl_strong_internal A P l H (prefix @ [hd]) tl ? ?) |
---|
64 | [ @ (H prefix hd tl prf acc) |
---|
65 | | applyS prf |
---|
66 | ] |
---|
67 | ] |
---|
68 | qed. |
---|
69 | |
---|
70 | definition foldl_strong ≝ |
---|
71 | λA: Type[0]. |
---|
72 | λP: list A → Type[0]. |
---|
73 | λl: list A. |
---|
74 | λH: ∀prefix. ∀hd. ∀tl. l = prefix @ [hd] @ tl → P prefix → P (prefix @ [hd]). |
---|
75 | λacc: P [ ]. |
---|
76 | foldl_strong_internal A P l H [ ] l acc (refl …). |
---|
77 | |
---|
78 | lemma pair_destruct: ∀A,B,a1,a2,b1,b2. pair A B a1 a2 = 〈b1,b2〉 → a1=b1 ∧ a2=b2. |
---|
79 | #A #B #a1 #a2 #b1 #b2 #EQ destruct /2/ |
---|
80 | qed. |
---|